Primitive. The name Eps_i is a term of type (setprop)set.
Axiom. (Eps_i_ax) We take the following as an axiom:
∀P : setprop, ∀x : set, P xP (Eps_i P)
Definition. We define True to be ∀p : prop, pp of type prop.
Definition. We define False to be ∀p : prop, p of type prop.
Definition. We define not to be λA : propAFalse of type propprop.
Notation. We use ¬ as a prefix operator with priority 700 corresponding to applying term not.
Definition. We define and to be λA B : prop∀p : prop, (ABp)p of type proppropprop.
Notation. We use as an infix operator with priority 780 and which associates to the left corresponding to applying term and.
Definition. We define or to be λA B : prop∀p : prop, (Ap)(Bp)p of type proppropprop.
Notation. We use as an infix operator with priority 785 and which associates to the left corresponding to applying term or.
Definition. We define iff to be λA B : propand (AB) (BA) of type proppropprop.
Notation. We use as an infix operator with priority 805 and no associativity corresponding to applying term iff.
Beginning of Section Eq
Variable A : SType
Definition. We define eq to be λx y : A∀Q : AAprop, Q x yQ y x of type AAprop.
Definition. We define neq to be λx y : A¬ eq x y of type AAprop.
End of Section Eq
Notation. We use = as an infix operator with priority 502 and no associativity corresponding to applying term eq.
Notation. We use as an infix operator with priority 502 and no associativity corresponding to applying term neq.
Beginning of Section FE
Variable A B : SType
Axiom. (func_ext) We take the following as an axiom:
∀f g : AB, (∀x : A, f x = g x)f = g
End of Section FE
Beginning of Section Ex
Variable A : SType
Definition. We define ex to be λQ : Aprop∀P : prop, (∀x : A, Q xP)P of type (Aprop)prop.
End of Section Ex
Notation. We use x...y [possibly with ascriptions] , B as a binder notation corresponding to a term constructed using ex.
Axiom. (prop_ext) We take the following as an axiom:
∀p q : prop, iff p qp = q
Primitive. The name In is a term of type setsetprop.
Notation. We use as an infix operator with priority 500 and no associativity corresponding to applying term In. Furthermore, we may write xA, B to mean x : set, xAB.
Definition. We define Subq to be λA B ⇒ xA, x B of type setsetprop.
Notation. We use as an infix operator with priority 500 and no associativity corresponding to applying term Subq. Furthermore, we may write xA, B to mean x : set, xAB.
Axiom. (set_ext) We take the following as an axiom:
∀X Y : set, X YY XX = Y
Axiom. (In_ind) We take the following as an axiom:
∀P : setprop, (∀X : set, (xX, P x)P X)∀X : set, P X
Notation. We use x...y [possibly with ascriptions] , B as a binder notation corresponding to a term constructed using ex and handling ∈ or ⊆ ascriptions using and.
Primitive. The name Empty is a term of type set.
Axiom. (EmptyAx) We take the following as an axiom:
¬ x : set, x Empty
Primitive. The name is a term of type setset.
Axiom. (UnionEq) We take the following as an axiom:
∀X x, x X Y, x Y Y X
Primitive. The name 𝒫 is a term of type setset.
Axiom. (PowerEq) We take the following as an axiom:
∀X Y : set, Y 𝒫 X Y X
Primitive. The name Repl is a term of type set(setset)set.
Notation. {B| xA} is notation for Repl Ax . B).
Axiom. (ReplEq) We take the following as an axiom:
∀A : set, ∀F : setset, ∀y : set, y {F x|xA} xA, y = F x
Definition. We define TransSet to be λU : setxU, x U of type setprop.
Definition. We define Union_closed to be λU : set∀X : set, X U X U of type setprop.
Definition. We define Power_closed to be λU : set∀X : set, X U𝒫 X U of type setprop.
Definition. We define Repl_closed to be λU : set∀X : set, X U∀F : setset, (∀x : set, x XF x U){F x|xX} U of type setprop.
Definition. We define ZF_closed to be λU : setUnion_closed U Power_closed U Repl_closed U of type setprop.
Primitive. The name UnivOf is a term of type setset.
Axiom. (UnivOf_In) We take the following as an axiom:
∀N : set, N UnivOf N
Axiom. (UnivOf_TransSet) We take the following as an axiom:
∀N : set, TransSet (UnivOf N)
Axiom. (UnivOf_ZF_closed) We take the following as an axiom:
∀N : set, ZF_closed (UnivOf N)
Axiom. (UnivOf_Min) We take the following as an axiom:
∀N U : set, N UTransSet UZF_closed UUnivOf N U
Axiom. (FalseE) We take the following as an axiom:
False∀p : prop, p
Axiom. (TrueI) We take the following as an axiom:
True
Axiom. (andI) We take the following as an axiom:
∀A B : prop, ABA B
Axiom. (andEL) We take the following as an axiom:
∀A B : prop, A BA
Axiom. (andER) We take the following as an axiom:
∀A B : prop, A BB
Axiom. (orIL) We take the following as an axiom:
∀A B : prop, AA B
Axiom. (orIR) We take the following as an axiom:
∀A B : prop, BA B
Beginning of Section PropN
Variable P1 P2 P3 : prop
Axiom. (and3I) We take the following as an axiom:
P1P2P3P1 P2 P3
Axiom. (and3E) We take the following as an axiom:
P1 P2 P3(∀p : prop, (P1P2P3p)p)
Axiom. (or3I1) We take the following as an axiom:
P1P1 P2 P3
Axiom. (or3I2) We take the following as an axiom:
P2P1 P2 P3
Axiom. (or3I3) We take the following as an axiom:
P3P1 P2 P3
Axiom. (or3E) We take the following as an axiom:
P1 P2 P3(∀p : prop, (P1p)(P2p)(P3p)p)
Variable P4 : prop
Axiom. (and4I) We take the following as an axiom:
P1P2P3P4P1 P2 P3 P4
Variable P5 : prop
Axiom. (and5I) We take the following as an axiom:
P1P2P3P4P5P1 P2 P3 P4 P5
End of Section PropN
Axiom. (not_or_and_demorgan) We take the following as an axiom:
∀A B : prop, ¬ (A B)¬ A ¬ B
Axiom. (not_ex_all_demorgan_i) We take the following as an axiom:
∀P : setprop, (¬ x, P x)∀x, ¬ P x
Axiom. (iffI) We take the following as an axiom:
∀A B : prop, (AB)(BA)(A B)
Axiom. (iffEL) We take the following as an axiom:
∀A B : prop, (A B)AB
Axiom. (iffER) We take the following as an axiom:
∀A B : prop, (A B)BA
Axiom. (iff_refl) We take the following as an axiom:
∀A : prop, A A
Axiom. (iff_sym) We take the following as an axiom:
∀A B : prop, (A B)(B A)
Axiom. (iff_trans) We take the following as an axiom:
∀A B C : prop, (A B)(B C)(A C)
Axiom. (eq_i_tra) We take the following as an axiom:
∀x y z, x = yy = zx = z
Axiom. (f_eq_i) We take the following as an axiom:
∀f : setset, ∀x y, x = yf x = f y
Axiom. (neq_i_sym) We take the following as an axiom:
∀x y, x yy x
Definition. We define nIn to be λx X ⇒ ¬ In x X of type setsetprop.
Notation. We use as an infix operator with priority 502 and no associativity corresponding to applying term nIn.
Axiom. (Eps_i_ex) We take the following as an axiom:
∀P : setprop, (x, P x)P (Eps_i P)
Axiom. (pred_ext) We take the following as an axiom:
∀P Q : setprop, (∀x, P x Q x)P = Q
Axiom. (prop_ext_2) We take the following as an axiom:
∀p q : prop, (pq)(qp)p = q
Axiom. (Subq_ref) We take the following as an axiom:
∀X : set, X X
Axiom. (Subq_tra) We take the following as an axiom:
∀X Y Z : set, X YY ZX Z
Axiom. (Subq_contra) We take the following as an axiom:
∀X Y z : set, X Yz Yz X
Axiom. (EmptyE) We take the following as an axiom:
∀x : set, x Empty
Axiom. (Subq_Empty) We take the following as an axiom:
∀X : set, Empty X
Axiom. (Empty_Subq_eq) We take the following as an axiom:
∀X : set, X EmptyX = Empty
Axiom. (Empty_eq) We take the following as an axiom:
∀X : set, (∀x, x X)X = Empty
Axiom. (UnionI) We take the following as an axiom:
∀X x Y : set, x YY Xx X
Axiom. (UnionE) We take the following as an axiom:
∀X x : set, x XY : set, x Y Y X
Axiom. (UnionE_impred) We take the following as an axiom:
∀X x : set, x X∀p : prop, (∀Y : set, x YY Xp)p
Axiom. (PowerI) We take the following as an axiom:
∀X Y : set, Y XY 𝒫 X
Axiom. (PowerE) We take the following as an axiom:
∀X Y : set, Y 𝒫 XY X
Axiom. (Empty_In_Power) We take the following as an axiom:
∀X : set, Empty 𝒫 X
Axiom. (Self_In_Power) We take the following as an axiom:
∀X : set, X 𝒫 X
Axiom. (xm) We take the following as an axiom:
∀P : prop, P ¬ P
Axiom. (dneg) We take the following as an axiom:
∀P : prop, ¬ ¬ PP
Axiom. (not_all_ex_demorgan_i) We take the following as an axiom:
∀P : setprop, ¬ (∀x, P x)x, ¬ P x
Axiom. (eq_or_nand) We take the following as an axiom:
or = (λx y : prop¬ (¬ x ¬ y))
Primitive. The name exactly1of2 is a term of type proppropprop.
Axiom. (exactly1of2_I1) We take the following as an axiom:
∀A B : prop, A¬ Bexactly1of2 A B
Axiom. (exactly1of2_I2) We take the following as an axiom:
∀A B : prop, ¬ ABexactly1of2 A B
Axiom. (exactly1of2_E) We take the following as an axiom:
∀A B : prop, exactly1of2 A B∀p : prop, (A¬ Bp)(¬ ABp)p
Axiom. (exactly1of2_or) We take the following as an axiom:
∀A B : prop, exactly1of2 A BA B
Axiom. (ReplI) We take the following as an axiom:
∀A : set, ∀F : setset, ∀x : set, x AF x {F x|xA}
Axiom. (ReplE) We take the following as an axiom:
∀A : set, ∀F : setset, ∀y : set, y {F x|xA}xA, y = F x
Axiom. (ReplE_impred) We take the following as an axiom:
∀A : set, ∀F : setset, ∀y : set, y {F x|xA}∀p : prop, (∀x : set, x Ay = F xp)p
Axiom. (ReplE') We take the following as an axiom:
∀X, ∀f : setset, ∀p : setprop, (xX, p (f x))y{f x|xX}, p y
Axiom. (Repl_Empty) We take the following as an axiom:
∀F : setset, {F x|xEmpty} = Empty
Axiom. (ReplEq_ext_sub) We take the following as an axiom:
∀X, ∀F G : setset, (xX, F x = G x){F x|xX} {G x|xX}
Axiom. (ReplEq_ext) We take the following as an axiom:
∀X, ∀F G : setset, (xX, F x = G x){F x|xX} = {G x|xX}
Axiom. (Repl_inv_eq) We take the following as an axiom:
∀P : setprop, ∀f g : setset, (∀x, P xg (f x) = x)∀X, (xX, P x){g y|y{f x|xX}} = X
Axiom. (Repl_invol_eq) We take the following as an axiom:
∀P : setprop, ∀f : setset, (∀x, P xf (f x) = x)∀X, (xX, P x){f y|y{f x|xX}} = X
Primitive. The name If_i is a term of type propsetsetset.
Notation. if cond then T else E is notation corresponding to If_i type cond T E where type is the inferred type of T.
Axiom. (If_i_correct) We take the following as an axiom:
∀p : prop, ∀x y : set, p (if p then x else y) = x ¬ p (if p then x else y) = y
Axiom. (If_i_0) We take the following as an axiom:
∀p : prop, ∀x y : set, ¬ p(if p then x else y) = y
Axiom. (If_i_1) We take the following as an axiom:
∀p : prop, ∀x y : set, p(if p then x else y) = x
Axiom. (If_i_or) We take the following as an axiom:
∀p : prop, ∀x y : set, (if p then x else y) = x (if p then x else y) = y
Primitive. The name UPair is a term of type setsetset.
Notation. {x,y} is notation for UPair x y.
Axiom. (UPairE) We take the following as an axiom:
∀x y z : set, x {y,z}x = y x = z
Axiom. (UPairI1) We take the following as an axiom:
∀y z : set, y {y,z}
Axiom. (UPairI2) We take the following as an axiom:
∀y z : set, z {y,z}
Primitive. The name Sing is a term of type setset.
Notation. {x} is notation for Sing x.
Axiom. (SingI) We take the following as an axiom:
∀x : set, x {x}
Axiom. (SingE) We take the following as an axiom:
∀x y : set, y {x}y = x
Primitive. The name binunion is a term of type setsetset.
Notation. We use as an infix operator with priority 345 and which associates to the left corresponding to applying term binunion.
Axiom. (binunionI1) We take the following as an axiom:
∀X Y z : set, z Xz X Y
Axiom. (binunionI2) We take the following as an axiom:
∀X Y z : set, z Yz X Y
Axiom. (binunionE) We take the following as an axiom:
∀X Y z : set, z X Yz X z Y
Axiom. (binunionE') We take the following as an axiom:
∀X Y z, ∀p : prop, (z Xp)(z Yp)(z X Yp)
Axiom. (binunion_asso) We take the following as an axiom:
∀X Y Z : set, X (Y Z) = (X Y) Z
Axiom. (binunion_com_Subq) We take the following as an axiom:
∀X Y : set, X Y Y X
Axiom. (binunion_com) We take the following as an axiom:
∀X Y : set, X Y = Y X
Axiom. (binunion_idl) We take the following as an axiom:
∀X : set, Empty X = X
Axiom. (binunion_idr) We take the following as an axiom:
∀X : set, X Empty = X
Axiom. (binunion_Subq_1) We take the following as an axiom:
∀X Y : set, X X Y
Axiom. (binunion_Subq_2) We take the following as an axiom:
∀X Y : set, Y X Y
Axiom. (binunion_Subq_min) We take the following as an axiom:
∀X Y Z : set, X ZY ZX Y Z
Axiom. (Subq_binunion_eq) We take the following as an axiom:
∀X Y, (X Y) = (X Y = Y)
Definition. We define SetAdjoin to be λX y ⇒ X {y} of type setsetset.
Notation. We now use the set enumeration notation {...,...,...} in general. If 0 elements are given, then Empty is used to form the corresponding term. If 1 element is given, then Sing is used to form the corresponding term. If 2 elements are given, then UPair is used to form the corresponding term. If more than elements are given, then SetAdjoin is used to reduce to the case with one fewer elements.
Primitive. The name famunion is a term of type set(setset)set.
Notation. We use x [possibly with ascriptions] , B as a binder notation corresponding to a term constructed using famunion.
Axiom. (famunionI) We take the following as an axiom:
∀X : set, ∀F : (setset), ∀x y : set, x Xy F xy xXF x
Axiom. (famunionE) We take the following as an axiom:
∀X : set, ∀F : (setset), ∀y : set, y (xXF x)xX, y F x
Axiom. (famunionE_impred) We take the following as an axiom:
∀X : set, ∀F : (setset), ∀y : set, y (xXF x)∀p : prop, (∀x, x Xy F xp)p
Axiom. (famunion_Empty) We take the following as an axiom:
∀F : setset, (xEmptyF x) = Empty
Beginning of Section SepSec
Variable X : set
Variable P : setprop
Let z : setEps_i (λz ⇒ z X P z)
Let F : setsetλx ⇒ if P x then x else z
Primitive. The name Sep is a term of type set.
End of Section SepSec
Notation. {xA | B} is notation for Sep Ax . B).
Axiom. (SepI) We take the following as an axiom:
∀X : set, ∀P : (setprop), ∀x : set, x XP xx {xX|P x}
Axiom. (SepE) We take the following as an axiom:
∀X : set, ∀P : (setprop), ∀x : set, x {xX|P x}x X P x
Axiom. (SepE1) We take the following as an axiom:
∀X : set, ∀P : (setprop), ∀x : set, x {xX|P x}x X
Axiom. (SepE2) We take the following as an axiom:
∀X : set, ∀P : (setprop), ∀x : set, x {xX|P x}P x
Axiom. (Sep_Subq) We take the following as an axiom:
∀X : set, ∀P : setprop, {xX|P x} X
Axiom. (Sep_In_Power) We take the following as an axiom:
∀X : set, ∀P : setprop, {xX|P x} 𝒫 X
Primitive. The name ReplSep is a term of type set(setprop)(setset)set.
Notation. {B| xA, C} is notation for ReplSep Ax . C) (λ x . B).
Axiom. (ReplSepI) We take the following as an axiom:
∀X : set, ∀P : setprop, ∀F : setset, ∀x : set, x XP xF x {F x|xX, P x}
Axiom. (ReplSepE) We take the following as an axiom:
∀X : set, ∀P : setprop, ∀F : setset, ∀y : set, y {F x|xX, P x}x : set, x X P x y = F x
Axiom. (ReplSepE_impred) We take the following as an axiom:
∀X : set, ∀P : setprop, ∀F : setset, ∀y : set, y {F x|xX, P x}∀p : prop, (xX, P xy = F xp)p
Primitive. The name binintersect is a term of type setsetset.
Notation. We use as an infix operator with priority 340 and which associates to the left corresponding to applying term binintersect.
Axiom. (binintersectI) We take the following as an axiom:
∀X Y z, z Xz Yz X Y
Axiom. (binintersectE) We take the following as an axiom:
∀X Y z, z X Yz X z Y
Axiom. (binintersectE1) We take the following as an axiom:
∀X Y z, z X Yz X
Axiom. (binintersectE2) We take the following as an axiom:
∀X Y z, z X Yz Y
Axiom. (binintersect_Subq_1) We take the following as an axiom:
∀X Y : set, X Y X
Axiom. (binintersect_Subq_2) We take the following as an axiom:
∀X Y : set, X Y Y
Axiom. (binintersect_Subq_eq_1) We take the following as an axiom:
∀X Y, X YX Y = X
Axiom. (binintersect_Subq_max) We take the following as an axiom:
∀X Y Z : set, Z XZ YZ X Y
Axiom. (binintersect_com_Subq) We take the following as an axiom:
∀X Y : set, X Y Y X
Axiom. (binintersect_com) We take the following as an axiom:
∀X Y : set, X Y = Y X
Primitive. The name setminus is a term of type setsetset.
Notation. We use as an infix operator with priority 350 and no associativity corresponding to applying term setminus.
Axiom. (setminusI) We take the following as an axiom:
∀X Y z, (z X)(z Y)z X Y
Axiom. (setminusE) We take the following as an axiom:
∀X Y z, (z X Y)z X z Y
Axiom. (setminusE1) We take the following as an axiom:
∀X Y z, (z X Y)z X
Axiom. (setminus_Subq) We take the following as an axiom:
∀X Y : set, X Y X
Axiom. (setminus_Subq_contra) We take the following as an axiom:
∀X Y Z : set, Z YX Y X Z
Axiom. (setminus_In_Power) We take the following as an axiom:
∀A U, A U 𝒫 A
Axiom. (In_irref) We take the following as an axiom:
∀x, x x
Axiom. (In_no2cycle) We take the following as an axiom:
∀x y, x yy xFalse
Primitive. The name ordsucc is a term of type setset.
Axiom. (ordsuccI1) We take the following as an axiom:
∀x : set, x ordsucc x
Axiom. (ordsuccI2) We take the following as an axiom:
∀x : set, x ordsucc x
Axiom. (ordsuccE) We take the following as an axiom:
∀x y : set, y ordsucc xy x y = x
Notation. Natural numbers 0,1,2,... are notation for the terms formed using Empty as 0 and forming successors with ordsucc.
Axiom. (neq_0_ordsucc) We take the following as an axiom:
∀a : set, 0 ordsucc a
Axiom. (neq_ordsucc_0) We take the following as an axiom:
∀a : set, ordsucc a 0
Axiom. (ordsucc_inj) We take the following as an axiom:
∀a b : set, ordsucc a = ordsucc ba = b
Axiom. (ordsucc_inj_contra) We take the following as an axiom:
∀a b : set, a bordsucc a ordsucc b
Axiom. (In_0_1) We take the following as an axiom:
0 1
Axiom. (In_0_2) We take the following as an axiom:
0 2
Axiom. (In_1_2) We take the following as an axiom:
1 2
Definition. We define nat_p to be λn : set∀p : setprop, p 0(∀x : set, p xp (ordsucc x))p n of type setprop.
Axiom. (nat_0) We take the following as an axiom:
nat_p 0
Axiom. (nat_ordsucc) We take the following as an axiom:
∀n : set, nat_p nnat_p (ordsucc n)
Axiom. (nat_1) We take the following as an axiom:
nat_p 1
Axiom. (nat_2) We take the following as an axiom:
nat_p 2
Axiom. (nat_0_in_ordsucc) We take the following as an axiom:
∀n, nat_p n0 ordsucc n
Axiom. (nat_ordsucc_in_ordsucc) We take the following as an axiom:
∀n, nat_p nmn, ordsucc m ordsucc n
Axiom. (nat_ind) We take the following as an axiom:
∀p : setprop, p 0(∀n, nat_p np np (ordsucc n))∀n, nat_p np n
Axiom. (nat_inv_impred) We take the following as an axiom:
∀p : setprop, p 0(∀n, nat_p np (ordsucc n))∀n, nat_p np n
Axiom. (nat_inv) We take the following as an axiom:
∀n, nat_p nn = 0 x, nat_p x n = ordsucc x
Axiom. (nat_complete_ind) We take the following as an axiom:
∀p : setprop, (∀n, nat_p n(mn, p m)p n)∀n, nat_p np n
Axiom. (nat_p_trans) We take the following as an axiom:
∀n, nat_p nmn, nat_p m
Axiom. (nat_trans) We take the following as an axiom:
∀n, nat_p nmn, m n
Axiom. (nat_ordsucc_trans) We take the following as an axiom:
∀n, nat_p nmordsucc n, m n
Axiom. (Union_ordsucc_eq) We take the following as an axiom:
∀n, nat_p n (ordsucc n) = n
Axiom. (cases_1) We take the following as an axiom:
i1, ∀p : setprop, p 0p i
Axiom. (cases_2) We take the following as an axiom:
i2, ∀p : setprop, p 0p 1p i
Axiom. (cases_3) We take the following as an axiom:
i3, ∀p : setprop, p 0p 1p 2p i
Axiom. (neq_0_1) We take the following as an axiom:
0 1
Axiom. (neq_1_0) We take the following as an axiom:
1 0
Axiom. (neq_0_2) We take the following as an axiom:
0 2
Axiom. (neq_2_0) We take the following as an axiom:
2 0
Axiom. (neq_1_2) We take the following as an axiom:
1 2
Axiom. (ZF_closed_E) We take the following as an axiom:
∀U, ZF_closed U∀p : prop, (Union_closed UPower_closed URepl_closed Up)p
Axiom. (ZF_Union_closed) We take the following as an axiom:
∀U, ZF_closed UXU, X U
Axiom. (ZF_Power_closed) We take the following as an axiom:
∀U, ZF_closed UXU, 𝒫 X U
Axiom. (ZF_Repl_closed) We take the following as an axiom:
∀U, ZF_closed UXU, ∀F : setset, (xX, F x U){F x|xX} U
Axiom. (ZF_UPair_closed) We take the following as an axiom:
∀U, ZF_closed Ux yU, {x,y} U
Axiom. (ZF_Sing_closed) We take the following as an axiom:
∀U, ZF_closed UxU, {x} U
Axiom. (ZF_binunion_closed) We take the following as an axiom:
∀U, ZF_closed UX YU, (X Y) U
Axiom. (ZF_ordsucc_closed) We take the following as an axiom:
∀U, ZF_closed UxU, ordsucc x U
Axiom. (nat_p_UnivOf_Empty) We take the following as an axiom:
∀n : set, nat_p nn UnivOf Empty
Primitive. The name ω is a term of type set.
Axiom. (omega_nat_p) We take the following as an axiom:
nω, nat_p n
Axiom. (nat_p_omega) We take the following as an axiom:
∀n : set, nat_p nn ω
Axiom. (omega_ordsucc) We take the following as an axiom:
nω, ordsucc n ω
Definition. We define ordinal to be λalpha : setTransSet alpha betaalpha, TransSet beta of type setprop.
Axiom. (ordinal_TransSet) We take the following as an axiom:
∀alpha : set, ordinal alphaTransSet alpha
Axiom. (ordinal_Empty) We take the following as an axiom:
ordinal Empty
Axiom. (ordinal_Hered) We take the following as an axiom:
∀alpha : set, ordinal alphabetaalpha, ordinal beta
Axiom. (TransSet_ordsucc) We take the following as an axiom:
∀X : set, TransSet XTransSet (ordsucc X)
Axiom. (ordinal_ordsucc) We take the following as an axiom:
∀alpha : set, ordinal alphaordinal (ordsucc alpha)
Axiom. (nat_p_ordinal) We take the following as an axiom:
∀n : set, nat_p nordinal n
Axiom. (ordinal_1) We take the following as an axiom:
ordinal 1
Axiom. (ordinal_2) We take the following as an axiom:
ordinal 2
Axiom. (omega_TransSet) We take the following as an axiom:
TransSet ω
Axiom. (omega_ordinal) We take the following as an axiom:
ordinal ω
Axiom. (ordsucc_omega_ordinal) We take the following as an axiom:
ordinal (ordsucc ω)
Axiom. (TransSet_ordsucc_In_Subq) We take the following as an axiom:
∀X : set, TransSet XxX, ordsucc x X
Axiom. (ordinal_ordsucc_In_Subq) We take the following as an axiom:
∀alpha, ordinal alphabetaalpha, ordsucc beta alpha
Axiom. (ordinal_trichotomy_or) We take the following as an axiom:
∀alpha beta : set, ordinal alphaordinal betaalpha beta alpha = beta beta alpha
Axiom. (ordinal_trichotomy_or_impred) We take the following as an axiom:
∀alpha beta : set, ordinal alphaordinal beta∀p : prop, (alpha betap)(alpha = betap)(beta alphap)p
Axiom. (ordinal_In_Or_Subq) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaalpha beta beta alpha
Axiom. (ordinal_linear) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaalpha beta beta alpha
Axiom. (ordinal_ordsucc_In_eq) We take the following as an axiom:
∀alpha beta, ordinal alphabeta alphaordsucc beta alpha alpha = ordsucc beta
Axiom. (ordinal_lim_or_succ) We take the following as an axiom:
∀alpha, ordinal alpha(betaalpha, ordsucc beta alpha) (betaalpha, alpha = ordsucc beta)
Axiom. (ordinal_ordsucc_In) We take the following as an axiom:
∀alpha, ordinal alphabetaalpha, ordsucc beta ordsucc alpha
Axiom. (ordinal_famunion) We take the following as an axiom:
∀X, ∀F : setset, (xX, ordinal (F x))ordinal (xXF x)
Axiom. (ordinal_binintersect) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaordinal (alpha beta)
Axiom. (ordinal_binunion) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaordinal (alpha beta)
Axiom. (ordinal_ind) We take the following as an axiom:
∀p : setprop, (∀alpha, ordinal alpha(betaalpha, p beta)p alpha)∀alpha, ordinal alphap alpha
Axiom. (least_ordinal_ex) We take the following as an axiom:
∀p : setprop, (alpha, ordinal alpha p alpha)alpha, ordinal alpha p alpha betaalpha, ¬ p beta
Definition. We define inj to be λX Y f ⇒ (uX, f u Y) (u vX, f u = f vu = v) of type setset(setset)prop.
Definition. We define bij to be λX Y f ⇒ (uX, f u Y) (u vX, f u = f vu = v) (wY, uX, f u = w) of type setset(setset)prop.
Axiom. (bijI) We take the following as an axiom:
∀X Y, ∀f : setset, (uX, f u Y)(u vX, f u = f vu = v)(wY, uX, f u = w)bij X Y f
Axiom. (bijE) We take the following as an axiom:
∀X Y, ∀f : setset, bij X Y f∀p : prop, ((uX, f u Y)(u vX, f u = f vu = v)(wY, uX, f u = w)p)p
Primitive. The name inv is a term of type set(setset)setset.
Axiom. (surj_rinv) We take the following as an axiom:
∀X Y, ∀f : setset, (wY, uX, f u = w)yY, inv X f y X f (inv X f y) = y
Axiom. (inj_linv) We take the following as an axiom:
∀X, ∀f : setset, (u vX, f u = f vu = v)xX, inv X f (f x) = x
Axiom. (bij_inv) We take the following as an axiom:
∀X Y, ∀f : setset, bij X Y fbij Y X (inv X f)
Axiom. (bij_id) We take the following as an axiom:
∀X, bij X X (λx ⇒ x)
Axiom. (bij_comp) We take the following as an axiom:
∀X Y Z : set, ∀f g : setset, bij X Y fbij Y Z gbij X Z (λx ⇒ g (f x))
Definition. We define equip to be λX Y : setf : setset, bij X Y f of type setsetprop.
Axiom. (equip_ref) We take the following as an axiom:
∀X, equip X X
Axiom. (equip_sym) We take the following as an axiom:
∀X Y, equip X Yequip Y X
Axiom. (equip_tra) We take the following as an axiom:
∀X Y Z, equip X Yequip Y Zequip X Z
Axiom. (equip_0_Empty) We take the following as an axiom:
∀X, equip X 0X = 0
Beginning of Section SchroederBernstein
Axiom. (KnasterTarski_set) We take the following as an axiom:
∀A, ∀F : setset, (U𝒫 A, F U 𝒫 A)(U V𝒫 A, U VF U F V)Y𝒫 A, F Y = Y
Axiom. (image_In_Power) We take the following as an axiom:
∀A B, ∀f : setset, (xA, f x B)U𝒫 A, {f x|xU} 𝒫 B
Axiom. (image_monotone) We take the following as an axiom:
∀f : setset, ∀U V, U V{f x|xU} {f x|xV}
Axiom. (setminus_antimonotone) We take the following as an axiom:
∀A U V, U VA V A U
Axiom. (SchroederBernstein) We take the following as an axiom:
∀A B, ∀f g : setset, inj A B finj B A gequip A B
End of Section SchroederBernstein
Beginning of Section PigeonHole
Axiom. (PigeonHole_nat) We take the following as an axiom:
∀n, nat_p n∀f : setset, (iordsucc n, f i n)¬ (i jordsucc n, f i = f ji = j)
Axiom. (PigeonHole_nat_bij) We take the following as an axiom:
∀n, nat_p n∀f : setset, (in, f i n)(i jn, f i = f ji = j)bij n n f
End of Section PigeonHole
Definition. We define finite to be λX ⇒ nω, equip X n of type setprop.
Axiom. (finite_ind) We take the following as an axiom:
∀p : setprop, p Empty(∀X y, finite Xy Xp Xp (X {y}))∀X, finite Xp X
Axiom. (finite_Empty) We take the following as an axiom:
finite 0
Axiom. (adjoin_finite) We take the following as an axiom:
∀X y, finite Xfinite (X {y})
Axiom. (binunion_finite) We take the following as an axiom:
∀X, finite X∀Y, finite Yfinite (X Y)
Axiom. (famunion_nat_finite) We take the following as an axiom:
∀X : setset, ∀n, nat_p n(in, finite (X i))finite (inX i)
Axiom. (Subq_finite) We take the following as an axiom:
∀X, finite X∀Y, Y Xfinite Y
Axiom. (TransSet_In_ordsucc_Subq) We take the following as an axiom:
∀x y, TransSet yx ordsucc yx y
Axiom. (exandE_i) We take the following as an axiom:
∀P Q : setprop, (x, P x Q x)∀r : prop, (∀x, P xQ xr)r
Axiom. (exandE_ii) We take the following as an axiom:
∀P Q : (setset)prop, (x : setset, P x Q x)∀p : prop, (∀x : setset, P xQ xp)p
Axiom. (exandE_iii) We take the following as an axiom:
∀P Q : (setsetset)prop, (x : setsetset, P x Q x)∀p : prop, (∀x : setsetset, P xQ xp)p
Axiom. (exandE_iiii) We take the following as an axiom:
∀P Q : (setsetsetset)prop, (x : setsetsetset, P x Q x)∀p : prop, (∀x : setsetsetset, P xQ xp)p
Beginning of Section Descr_ii
Variable P : (setset)prop
Primitive. The name Descr_ii is a term of type setset.
Hypothesis Pex : f : setset, P f
Hypothesis Puniq : ∀f g : setset, P fP gf = g
Axiom. (Descr_ii_prop) We take the following as an axiom:
End of Section Descr_ii
Beginning of Section Descr_iii
Variable P : (setsetset)prop
Primitive. The name Descr_iii is a term of type setsetset.
Hypothesis Pex : f : setsetset, P f
Hypothesis Puniq : ∀f g : setsetset, P fP gf = g
Axiom. (Descr_iii_prop) We take the following as an axiom:
End of Section Descr_iii
Beginning of Section Descr_Vo1
Variable P : Vo 1prop
Primitive. The name Descr_Vo1 is a term of type Vo 1.
Hypothesis Pex : f : Vo 1, P f
Hypothesis Puniq : ∀f g : Vo 1, P fP gf = g
Axiom. (Descr_Vo1_prop) We take the following as an axiom:
End of Section Descr_Vo1
Beginning of Section If_ii
Variable p : prop
Variable f g : setset
Primitive. The name If_ii is a term of type setset.
Axiom. (If_ii_1) We take the following as an axiom:
pIf_ii = f
Axiom. (If_ii_0) We take the following as an axiom:
¬ pIf_ii = g
End of Section If_ii
Beginning of Section If_iii
Variable p : prop
Variable f g : setsetset
Primitive. The name If_iii is a term of type setsetset.
Axiom. (If_iii_1) We take the following as an axiom:
pIf_iii = f
Axiom. (If_iii_0) We take the following as an axiom:
¬ pIf_iii = g
End of Section If_iii
Beginning of Section EpsilonRec_i
Variable F : set(setset)set
Primitive. The name In_rec_i is a term of type setset.
Hypothesis Fr : ∀X : set, ∀g h : setset, (xX, g x = h x)F X g = F X h
Axiom. (In_rec_i_eq) We take the following as an axiom:
∀X : set, In_rec_i X = F X In_rec_i
End of Section EpsilonRec_i
Beginning of Section EpsilonRec_ii
Variable F : set(set(setset))(setset)
Primitive. The name In_rec_ii is a term of type set(setset).
Hypothesis Fr : ∀X : set, ∀g h : set(setset), (xX, g x = h x)F X g = F X h
Axiom. (In_rec_ii_eq) We take the following as an axiom:
∀X : set, In_rec_ii X = F X In_rec_ii
End of Section EpsilonRec_ii
Beginning of Section EpsilonRec_iii
Variable F : set(set(setsetset))(setsetset)
Primitive. The name In_rec_iii is a term of type set(setsetset).
Hypothesis Fr : ∀X : set, ∀g h : set(setsetset), (xX, g x = h x)F X g = F X h
Axiom. (In_rec_iii_eq) We take the following as an axiom:
∀X : set, In_rec_iii X = F X In_rec_iii
End of Section EpsilonRec_iii
Beginning of Section NatRec
Variable z : set
Variable f : setsetset
Let F : set(setset)setλn g ⇒ if n n then f ( n) (g ( n)) else z
Definition. We define nat_primrec to be In_rec_i F of type setset.
Axiom. (nat_primrec_r) We take the following as an axiom:
∀X : set, ∀g h : setset, (xX, g x = h x)F X g = F X h
Axiom. (nat_primrec_0) We take the following as an axiom:
Axiom. (nat_primrec_S) We take the following as an axiom:
∀n : set, nat_p nnat_primrec (ordsucc n) = f n (nat_primrec n)
End of Section NatRec
Beginning of Section NatArith
Definition. We define add_nat to be λn m : setnat_primrec n (λ_ r ⇒ ordsucc r) m of type setsetset.
Notation. We use + as an infix operator with priority 360 and which associates to the right corresponding to applying term add_nat.
Axiom. (add_nat_0R) We take the following as an axiom:
∀n : set, n + 0 = n
Axiom. (add_nat_SR) We take the following as an axiom:
∀n m : set, nat_p mn + ordsucc m = ordsucc (n + m)
Axiom. (add_nat_p) We take the following as an axiom:
∀n : set, nat_p n∀m : set, nat_p mnat_p (n + m)
Axiom. (add_nat_1_1_2) We take the following as an axiom:
1 + 1 = 2
Definition. We define mul_nat to be λn m : setnat_primrec 0 (λ_ r ⇒ n + r) m of type setsetset.
Notation. We use * as an infix operator with priority 355 and which associates to the right corresponding to applying term mul_nat.
Axiom. (mul_nat_0R) We take the following as an axiom:
∀n : set, n * 0 = 0
Axiom. (mul_nat_SR) We take the following as an axiom:
∀n m : set, nat_p mn * ordsucc m = n + n * m
Axiom. (mul_nat_p) We take the following as an axiom:
∀n : set, nat_p n∀m : set, nat_p mnat_p (n * m)
End of Section NatArith
Definition. We define Inj1 to be In_rec_i (λX f ⇒ {0} {f x|xX}) of type setset.
Axiom. (Inj1_eq) We take the following as an axiom:
∀X : set, Inj1 X = {0} {Inj1 x|xX}
Axiom. (Inj1I1) We take the following as an axiom:
∀X : set, 0 Inj1 X
Axiom. (Inj1I2) We take the following as an axiom:
∀X x : set, x XInj1 x Inj1 X
Axiom. (Inj1E) We take the following as an axiom:
∀X y : set, y Inj1 Xy = 0 xX, y = Inj1 x
Axiom. (Inj1NE1) We take the following as an axiom:
∀x : set, Inj1 x 0
Axiom. (Inj1NE2) We take the following as an axiom:
∀x : set, Inj1 x {0}
Definition. We define Inj0 to be λX ⇒ {Inj1 x|xX} of type setset.
Axiom. (Inj0I) We take the following as an axiom:
∀X x : set, x XInj1 x Inj0 X
Axiom. (Inj0E) We take the following as an axiom:
∀X y : set, y Inj0 Xx : set, x X y = Inj1 x
Definition. We define Unj to be In_rec_i (λX f ⇒ {f x|xX {0}}) of type setset.
Axiom. (Unj_eq) We take the following as an axiom:
∀X : set, Unj X = {Unj x|xX {0}}
Axiom. (Unj_Inj1_eq) We take the following as an axiom:
∀X : set, Unj (Inj1 X) = X
Axiom. (Inj1_inj) We take the following as an axiom:
∀X Y : set, Inj1 X = Inj1 YX = Y
Axiom. (Unj_Inj0_eq) We take the following as an axiom:
∀X : set, Unj (Inj0 X) = X
Axiom. (Inj0_inj) We take the following as an axiom:
∀X Y : set, Inj0 X = Inj0 YX = Y
Axiom. (Inj0_0) We take the following as an axiom:
Inj0 0 = 0
Axiom. (Inj0_Inj1_neq) We take the following as an axiom:
∀X Y : set, Inj0 X Inj1 Y
Definition. We define setsum to be λX Y ⇒ {Inj0 x|xX} {Inj1 y|yY} of type setsetset.
Notation. We use + as an infix operator with priority 450 and which associates to the left corresponding to applying term setsum.
Axiom. (Inj0_setsum) We take the following as an axiom:
∀X Y x : set, x XInj0 x X + Y
Axiom. (Inj1_setsum) We take the following as an axiom:
∀X Y y : set, y YInj1 y X + Y
Axiom. (setsum_Inj_inv) We take the following as an axiom:
∀X Y z : set, z X + Y(xX, z = Inj0 x) (yY, z = Inj1 y)
Axiom. (Inj0_setsum_0L) We take the following as an axiom:
∀X : set, 0 + X = Inj0 X
Axiom. (Subq_1_Sing0) We take the following as an axiom:
1 {0}
Axiom. (Inj1_setsum_1L) We take the following as an axiom:
∀X : set, 1 + X = Inj1 X
Axiom. (nat_setsum1_ordsucc) We take the following as an axiom:
∀n : set, nat_p n1 + n = ordsucc n
Axiom. (setsum_0_0) We take the following as an axiom:
0 + 0 = 0
Axiom. (setsum_1_0_1) We take the following as an axiom:
1 + 0 = 1
Axiom. (setsum_1_1_2) We take the following as an axiom:
1 + 1 = 2
Beginning of Section pair_setsum
Let pair ≝ setsum
Definition. We define proj0 to be λZ ⇒ {Unj z|zZ, x : set, Inj0 x = z} of type setset.
Definition. We define proj1 to be λZ ⇒ {Unj z|zZ, y : set, Inj1 y = z} of type setset.
Axiom. (Inj0_pair_0_eq) We take the following as an axiom:
Inj0 = pair 0
Axiom. (Inj1_pair_1_eq) We take the following as an axiom:
Inj1 = pair 1
Axiom. (pairI0) We take the following as an axiom:
∀X Y x, x Xpair 0 x pair X Y
Axiom. (pairI1) We take the following as an axiom:
∀X Y y, y Ypair 1 y pair X Y
Axiom. (pairE) We take the following as an axiom:
∀X Y z, z pair X Y(xX, z = pair 0 x) (yY, z = pair 1 y)
Axiom. (pairE0) We take the following as an axiom:
∀X Y x, pair 0 x pair X Yx X
Axiom. (pairE1) We take the following as an axiom:
∀X Y y, pair 1 y pair X Yy Y
Axiom. (proj0I) We take the following as an axiom:
∀w u : set, pair 0 u wu proj0 w
Axiom. (proj0E) We take the following as an axiom:
∀w u : set, u proj0 wpair 0 u w
Axiom. (proj1I) We take the following as an axiom:
∀w u : set, pair 1 u wu proj1 w
Axiom. (proj1E) We take the following as an axiom:
∀w u : set, u proj1 wpair 1 u w
Axiom. (proj0_pair_eq) We take the following as an axiom:
∀X Y : set, proj0 (pair X Y) = X
Axiom. (proj1_pair_eq) We take the following as an axiom:
∀X Y : set, proj1 (pair X Y) = Y
Definition. We define Sigma to be λX Y ⇒ xX{pair x y|yY x} of type set(setset)set.
Notation. We use x...y [possibly with ascriptions] , B as a binder notation corresponding to a term constructed using Sigma.
Axiom. (pair_Sigma) We take the following as an axiom:
∀X : set, ∀Y : setset, xX, yY x, pair x y xX, Y x
Axiom. (Sigma_eta_proj0_proj1) We take the following as an axiom:
∀X : set, ∀Y : setset, z(xX, Y x), pair (proj0 z) (proj1 z) = z proj0 z X proj1 z Y (proj0 z)
Axiom. (proj_Sigma_eta) We take the following as an axiom:
∀X : set, ∀Y : setset, z(xX, Y x), pair (proj0 z) (proj1 z) = z
Axiom. (proj0_Sigma) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀z : set, z (xX, Y x)proj0 z X
Axiom. (proj1_Sigma) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀z : set, z (xX, Y x)proj1 z Y (proj0 z)
Axiom. (pair_Sigma_E1) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀x y : set, pair x y (xX, Y x)y Y x
Axiom. (Sigma_E) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀z : set, z (xX, Y x)xX, yY x, z = pair x y
Definition. We define setprod to be λX Y : setxX, Y of type setsetset.
Notation. We use as an infix operator with priority 440 and which associates to the left corresponding to applying term setprod.
Let lam : set(setset)setSigma
Definition. We define ap to be λf x ⇒ {proj1 z|zf, y : set, z = pair x y} of type setsetset.
Notation. When x is a set, a term x y is notation for ap x y.
Notation. λ xAB is notation for the set Sigma Ax : set ⇒ B).
Notation. We now use n-tuple notation (a0,...,an-1) for n ≥ 2 for λ i ∈ n . if i = 0 then a0 else ... if i = n-2 then an-2 else an-1.
Axiom. (lamI) We take the following as an axiom:
∀X : set, ∀F : setset, xX, yF x, pair x y λxX F x
Axiom. (lamE) We take the following as an axiom:
∀X : set, ∀F : setset, ∀z : set, z (λxX F x)xX, yF x, z = pair x y
Axiom. (apI) We take the following as an axiom:
∀f x y, pair x y fy f x
Axiom. (apE) We take the following as an axiom:
∀f x y, y f xpair x y f
Axiom. (beta) We take the following as an axiom:
∀X : set, ∀F : setset, ∀x : set, x X(λxX F x) x = F x
Axiom. (proj0_ap_0) We take the following as an axiom:
∀u, proj0 u = u 0
Axiom. (proj1_ap_1) We take the following as an axiom:
∀u, proj1 u = u 1
Axiom. (pair_ap_0) We take the following as an axiom:
∀x y : set, (pair x y) 0 = x
Axiom. (pair_ap_1) We take the following as an axiom:
∀x y : set, (pair x y) 1 = y
Axiom. (ap0_Sigma) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀z : set, z (xX, Y x)(z 0) X
Axiom. (ap1_Sigma) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀z : set, z (xX, Y x)(z 1) (Y (z 0))
Definition. We define pair_p to be λu : setpair (u 0) (u 1) = u of type setprop.
Axiom. (pair_p_I) We take the following as an axiom:
∀x y, pair_p (pair x y)
Axiom. (Subq_2_UPair01) We take the following as an axiom:
Axiom. (tuple_pair) We take the following as an axiom:
∀x y : set, pair x y = (x,y)
Definition. We define Pi to be λX Y ⇒ {f𝒫 (xX, (Y x))|xX, f x Y x} of type set(setset)set.
Notation. We use x...y [possibly with ascriptions] , B as a binder notation corresponding to a term constructed using Pi.
Axiom. (PiI) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀f : set, (uf, pair_p u u 0 X)(xX, f x Y x)f xX, Y x
Axiom. (lam_Pi) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀F : setset, (xX, F x Y x)(λxX F x) (xX, Y x)
Axiom. (ap_Pi) We take the following as an axiom:
∀X : set, ∀Y : setset, ∀f : set, ∀x : set, f (xX, Y x)x Xf x Y x
Definition. We define setexp to be λX Y : setyY, X of type setsetset.
Notation. We use :^: as an infix operator with priority 430 and which associates to the left corresponding to applying term setexp.
Axiom. (pair_tuple_fun) We take the following as an axiom:
pair = (λx y ⇒ (x,y))
Axiom. (lamI2) We take the following as an axiom:
∀X, ∀F : setset, xX, yF x, (x,y) λxX F x
Beginning of Section Tuples
Variable x0 x1 : set
Axiom. (tuple_2_0_eq) We take the following as an axiom:
(x0,x1) 0 = x0
Axiom. (tuple_2_1_eq) We take the following as an axiom:
(x0,x1) 1 = x1
End of Section Tuples
Axiom. (ReplEq_setprod_ext) We take the following as an axiom:
∀X Y, ∀F G : setsetset, (xX, yY, F x y = G x y){F (w 0) (w 1)|wX Y} = {G (w 0) (w 1)|wX Y}
Axiom. (tuple_2_Sigma) We take the following as an axiom:
∀X : set, ∀Y : setset, xX, yY x, (x,y) xX, Y x
Axiom. (tuple_2_setprod) We take the following as an axiom:
∀X : set, ∀Y : set, xX, yY, (x,y) X Y
End of Section pair_setsum
Notation. We use x...y [possibly with ascriptions] , B as a binder notation corresponding to a term constructed using Sigma.
Notation. We use as an infix operator with priority 440 and which associates to the left corresponding to applying term setprod.
Notation. We use x...y [possibly with ascriptions] , B as a binder notation corresponding to a term constructed using Pi.
Notation. We use :^: as an infix operator with priority 430 and which associates to the left corresponding to applying term setexp.
Primitive. The name DescrR_i_io_1 is a term of type (set(setprop)prop)set.
Primitive. The name DescrR_i_io_2 is a term of type (set(setprop)prop)setprop.
Axiom. (DescrR_i_io_12) We take the following as an axiom:
∀R : set(setprop)prop, (x, (y : setprop, R x y) (∀y z : setprop, R x yR x zy = z))R (DescrR_i_io_1 R) (DescrR_i_io_2 R)
Definition. We define PNoEq_ to be λalpha p q ⇒ betaalpha, p beta q beta of type set(setprop)(setprop)prop.
Axiom. (PNoEq_ref_) We take the following as an axiom:
∀alpha, ∀p : setprop, PNoEq_ alpha p p
Axiom. (PNoEq_sym_) We take the following as an axiom:
∀alpha, ∀p q : setprop, PNoEq_ alpha p qPNoEq_ alpha q p
Axiom. (PNoEq_tra_) We take the following as an axiom:
∀alpha, ∀p q r : setprop, PNoEq_ alpha p qPNoEq_ alpha q rPNoEq_ alpha p r
Axiom. (PNoEq_antimon_) We take the following as an axiom:
∀p q : setprop, ∀alpha, ordinal alphabetaalpha, PNoEq_ alpha p qPNoEq_ beta p q
Definition. We define PNoLt_ to be λalpha p q ⇒ betaalpha, PNoEq_ beta p q ¬ p beta q beta of type set(setprop)(setprop)prop.
Axiom. (PNoLt_E_) We take the following as an axiom:
∀alpha, ∀p q : setprop, PNoLt_ alpha p q∀R : prop, (∀beta, beta alphaPNoEq_ beta p q¬ p betaq betaR)R
Axiom. (PNoLt_irref_) We take the following as an axiom:
∀alpha, ∀p : setprop, ¬ PNoLt_ alpha p p
Axiom. (PNoLt_mon_) We take the following as an axiom:
∀p q : setprop, ∀alpha, ordinal alphabetaalpha, PNoLt_ beta p qPNoLt_ alpha p q
Axiom. (PNoLt_trichotomy_or_) We take the following as an axiom:
∀p q : setprop, ∀alpha, ordinal alphaPNoLt_ alpha p q PNoEq_ alpha p q PNoLt_ alpha q p
Axiom. (PNoLt_tra_) We take the following as an axiom:
∀alpha, ordinal alpha∀p q r : setprop, PNoLt_ alpha p qPNoLt_ alpha q rPNoLt_ alpha p r
Primitive. The name PNoLt is a term of type set(setprop)set(setprop)prop.
Axiom. (PNoLtI1) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, PNoLt_ (alpha beta) p qPNoLt alpha p beta q
Axiom. (PNoLtI2) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, alpha betaPNoEq_ alpha p qq alphaPNoLt alpha p beta q
Axiom. (PNoLtI3) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, beta alphaPNoEq_ beta p q¬ p betaPNoLt alpha p beta q
Axiom. (PNoLtE) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, PNoLt alpha p beta q∀R : prop, (PNoLt_ (alpha beta) p qR)(alpha betaPNoEq_ alpha p qq alphaR)(beta alphaPNoEq_ beta p q¬ p betaR)R
Axiom. (PNoLt_irref) We take the following as an axiom:
∀alpha, ∀p : setprop, ¬ PNoLt alpha p alpha p
Axiom. (PNoLt_trichotomy_or) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, ordinal alphaordinal betaPNoLt alpha p beta q alpha = beta PNoEq_ alpha p q PNoLt beta q alpha p
Axiom. (PNoLtEq_tra) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal beta∀p q r : setprop, PNoLt alpha p beta qPNoEq_ beta q rPNoLt alpha p beta r
Axiom. (PNoEqLt_tra) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal beta∀p q r : setprop, PNoEq_ alpha p qPNoLt alpha q beta rPNoLt alpha p beta r
Axiom. (PNoLt_tra) We take the following as an axiom:
∀alpha beta gamma, ordinal alphaordinal betaordinal gamma∀p q r : setprop, PNoLt alpha p beta qPNoLt beta q gamma rPNoLt alpha p gamma r
Definition. We define PNoLe to be λalpha p beta q ⇒ PNoLt alpha p beta q alpha = beta PNoEq_ alpha p q of type set(setprop)set(setprop)prop.
Axiom. (PNoLeI1) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, PNoLt alpha p beta qPNoLe alpha p beta q
Axiom. (PNoLeI2) We take the following as an axiom:
∀alpha, ∀p q : setprop, PNoEq_ alpha p qPNoLe alpha p alpha q
Axiom. (PNoLe_ref) We take the following as an axiom:
∀alpha, ∀p : setprop, PNoLe alpha p alpha p
Axiom. (PNoLe_antisym) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal beta∀p q : setprop, PNoLe alpha p beta qPNoLe beta q alpha palpha = beta PNoEq_ alpha p q
Axiom. (PNoLtLe_tra) We take the following as an axiom:
∀alpha beta gamma, ordinal alphaordinal betaordinal gamma∀p q r : setprop, PNoLt alpha p beta qPNoLe beta q gamma rPNoLt alpha p gamma r
Axiom. (PNoLeLt_tra) We take the following as an axiom:
∀alpha beta gamma, ordinal alphaordinal betaordinal gamma∀p q r : setprop, PNoLe alpha p beta qPNoLt beta q gamma rPNoLt alpha p gamma r
Axiom. (PNoEqLe_tra) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal beta∀p q r : setprop, PNoEq_ alpha p qPNoLe alpha q beta rPNoLe alpha p beta r
Axiom. (PNoLe_tra) We take the following as an axiom:
∀alpha beta gamma, ordinal alphaordinal betaordinal gamma∀p q r : setprop, PNoLe alpha p beta qPNoLe beta q gamma rPNoLe alpha p gamma r
Definition. We define PNo_downc to be λL alpha p ⇒ beta, ordinal beta q : setprop, L beta q PNoLe alpha p beta q of type (set(setprop)prop)set(setprop)prop.
Definition. We define PNo_upc to be λR alpha p ⇒ beta, ordinal beta q : setprop, R beta q PNoLe beta q alpha p of type (set(setprop)prop)set(setprop)prop.
Axiom. (PNoLe_downc) We take the following as an axiom:
∀L : set(setprop)prop, ∀alpha beta, ∀p q : setprop, ordinal alphaordinal betaPNo_downc L alpha pPNoLe beta q alpha pPNo_downc L beta q
Axiom. (PNo_downc_ref) We take the following as an axiom:
∀L : set(setprop)prop, ∀alpha, ordinal alpha∀p : setprop, L alpha pPNo_downc L alpha p
Axiom. (PNo_upc_ref) We take the following as an axiom:
∀R : set(setprop)prop, ∀alpha, ordinal alpha∀p : setprop, R alpha pPNo_upc R alpha p
Axiom. (PNoLe_upc) We take the following as an axiom:
∀R : set(setprop)prop, ∀alpha beta, ∀p q : setprop, ordinal alphaordinal betaPNo_upc R alpha pPNoLe alpha p beta qPNo_upc R beta q
Definition. We define PNoLt_pwise to be λL R ⇒ ∀gamma, ordinal gamma∀p : setprop, L gamma p∀delta, ordinal delta∀q : setprop, R delta qPNoLt gamma p delta q of type (set(setprop)prop)(set(setprop)prop)prop.
Axiom. (PNoLt_pwise_downc_upc) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L RPNoLt_pwise (PNo_downc L) (PNo_upc R)
Definition. We define PNo_rel_strict_upperbd to be λL alpha p ⇒ betaalpha, ∀q : setprop, PNo_downc L beta qPNoLt beta q alpha p of type (set(setprop)prop)set(setprop)prop.
Definition. We define PNo_rel_strict_lowerbd to be λR alpha p ⇒ betaalpha, ∀q : setprop, PNo_upc R beta qPNoLt alpha p beta q of type (set(setprop)prop)set(setprop)prop.
Definition. We define PNo_rel_strict_imv to be λL R alpha p ⇒ PNo_rel_strict_upperbd L alpha p PNo_rel_strict_lowerbd R alpha p of type (set(setprop)prop)(set(setprop)prop)set(setprop)prop.
Axiom. (PNoEq_rel_strict_upperbd) We take the following as an axiom:
∀L : set(setprop)prop, ∀alpha, ordinal alpha∀p q : setprop, PNoEq_ alpha p qPNo_rel_strict_upperbd L alpha pPNo_rel_strict_upperbd L alpha q
Axiom. (PNo_rel_strict_upperbd_antimon) We take the following as an axiom:
∀L : set(setprop)prop, ∀alpha, ordinal alpha∀p : setprop, betaalpha, PNo_rel_strict_upperbd L alpha pPNo_rel_strict_upperbd L beta p
Axiom. (PNoEq_rel_strict_lowerbd) We take the following as an axiom:
∀R : set(setprop)prop, ∀alpha, ordinal alpha∀p q : setprop, PNoEq_ alpha p qPNo_rel_strict_lowerbd R alpha pPNo_rel_strict_lowerbd R alpha q
Axiom. (PNo_rel_strict_lowerbd_antimon) We take the following as an axiom:
∀R : set(setprop)prop, ∀alpha, ordinal alpha∀p : setprop, betaalpha, PNo_rel_strict_lowerbd R alpha pPNo_rel_strict_lowerbd R beta p
Axiom. (PNoEq_rel_strict_imv) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alpha∀p q : setprop, PNoEq_ alpha p qPNo_rel_strict_imv L R alpha pPNo_rel_strict_imv L R alpha q
Axiom. (PNo_rel_strict_imv_antimon) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alpha∀p : setprop, betaalpha, PNo_rel_strict_imv L R alpha pPNo_rel_strict_imv L R beta p
Definition. We define PNo_rel_strict_uniq_imv to be λL R alpha p ⇒ PNo_rel_strict_imv L R alpha p ∀q : setprop, PNo_rel_strict_imv L R alpha qPNoEq_ alpha p q of type (set(setprop)prop)(set(setprop)prop)set(setprop)prop.
Definition. We define PNo_rel_strict_split_imv to be λL R alpha p ⇒ PNo_rel_strict_imv L R (ordsucc alpha) (λdelta ⇒ p delta delta alpha) PNo_rel_strict_imv L R (ordsucc alpha) (λdelta ⇒ p delta delta = alpha) of type (set(setprop)prop)(set(setprop)prop)set(setprop)prop.
Axiom. (PNo_extend0_eq) We take the following as an axiom:
∀alpha, ∀p : setprop, PNoEq_ alpha p (λdelta ⇒ p delta delta alpha)
Axiom. (PNo_extend1_eq) We take the following as an axiom:
∀alpha, ∀p : setprop, PNoEq_ alpha p (λdelta ⇒ p delta delta = alpha)
Axiom. (PNo_rel_imv_ex) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alpha(p : setprop, PNo_rel_strict_uniq_imv L R alpha p) (taualpha, p : setprop, PNo_rel_strict_split_imv L R tau p)
Definition. We define PNo_lenbdd to be λalpha L ⇒ ∀beta, ∀p : setprop, L beta pbeta alpha of type set(set(setprop)prop)prop.
Axiom. (PNo_lenbdd_strict_imv_extend0) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha R∀p : setprop, PNo_rel_strict_imv L R alpha pPNo_rel_strict_imv L R (ordsucc alpha) (λdelta ⇒ p delta delta alpha)
Axiom. (PNo_lenbdd_strict_imv_extend1) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha R∀p : setprop, PNo_rel_strict_imv L R alpha pPNo_rel_strict_imv L R (ordsucc alpha) (λdelta ⇒ p delta delta = alpha)
Axiom. (PNo_lenbdd_strict_imv_split) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha R∀p : setprop, PNo_rel_strict_imv L R alpha pPNo_rel_strict_split_imv L R alpha p
Axiom. (PNo_rel_imv_bdd_ex) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha Rbetaordsucc alpha, p : setprop, PNo_rel_strict_split_imv L R beta p
Definition. We define PNo_strict_upperbd to be λL alpha p ⇒ ∀beta, ordinal beta∀q : setprop, L beta qPNoLt beta q alpha p of type (set(setprop)prop)set(setprop)prop.
Definition. We define PNo_strict_lowerbd to be λR alpha p ⇒ ∀beta, ordinal beta∀q : setprop, R beta qPNoLt alpha p beta q of type (set(setprop)prop)set(setprop)prop.
Definition. We define PNo_strict_imv to be λL R alpha p ⇒ PNo_strict_upperbd L alpha p PNo_strict_lowerbd R alpha p of type (set(setprop)prop)(set(setprop)prop)set(setprop)prop.
Axiom. (PNoEq_strict_upperbd) We take the following as an axiom:
∀L : set(setprop)prop, ∀alpha, ordinal alpha∀p q : setprop, PNoEq_ alpha p qPNo_strict_upperbd L alpha pPNo_strict_upperbd L alpha q
Axiom. (PNoEq_strict_lowerbd) We take the following as an axiom:
∀R : set(setprop)prop, ∀alpha, ordinal alpha∀p q : setprop, PNoEq_ alpha p qPNo_strict_lowerbd R alpha pPNo_strict_lowerbd R alpha q
Axiom. (PNoEq_strict_imv) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alpha∀p q : setprop, PNoEq_ alpha p qPNo_strict_imv L R alpha pPNo_strict_imv L R alpha q
Axiom. (PNo_strict_upperbd_imp_rel_strict_upperbd) We take the following as an axiom:
∀L : set(setprop)prop, ∀alpha, ordinal alphabetaordsucc alpha, ∀p : setprop, PNo_strict_upperbd L alpha pPNo_rel_strict_upperbd L beta p
Axiom. (PNo_strict_lowerbd_imp_rel_strict_lowerbd) We take the following as an axiom:
∀R : set(setprop)prop, ∀alpha, ordinal alphabetaordsucc alpha, ∀p : setprop, PNo_strict_lowerbd R alpha pPNo_rel_strict_lowerbd R beta p
Axiom. (PNo_strict_imv_imp_rel_strict_imv) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alphabetaordsucc alpha, ∀p : setprop, PNo_strict_imv L R alpha pPNo_rel_strict_imv L R beta p
Axiom. (PNo_rel_split_imv_imp_strict_imv) We take the following as an axiom:
∀L R : set(setprop)prop, ∀alpha, ordinal alpha∀p : setprop, PNo_rel_strict_split_imv L R alpha pPNo_strict_imv L R alpha p
Axiom. (PNo_lenbdd_strict_imv_ex) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha Rbetaordsucc alpha, p : setprop, PNo_strict_imv L R beta p
Definition. We define PNo_least_rep to be λL R beta p ⇒ ordinal beta PNo_strict_imv L R beta p gammabeta, ∀q : setprop, ¬ PNo_strict_imv L R gamma q of type (set(setprop)prop)(set(setprop)prop)set(setprop)prop.
Definition. We define PNo_least_rep2 to be λL R beta p ⇒ PNo_least_rep L R beta p ∀x, x beta¬ p x of type (set(setprop)prop)(set(setprop)prop)set(setprop)prop.
Axiom. (PNo_strict_imv_pred_eq) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alpha∀p q : setprop, PNo_least_rep L R alpha pPNo_strict_imv L R alpha qbetaalpha, p beta q beta
Axiom. (PNo_lenbdd_least_rep2_exuniq2) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha Rbeta, (p : setprop, PNo_least_rep2 L R beta p) (∀p q : setprop, PNo_least_rep2 L R beta pPNo_least_rep2 L R beta qp = q)
Primitive. The name PNo_bd is a term of type (set(setprop)prop)(set(setprop)prop)set.
Primitive. The name PNo_pred is a term of type (set(setprop)prop)(set(setprop)prop)setprop.
Axiom. (PNo_bd_pred_lem) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha RPNo_least_rep2 L R (PNo_bd L R) (PNo_pred L R)
Axiom. (PNo_bd_pred) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha RPNo_least_rep L R (PNo_bd L R) (PNo_pred L R)
Axiom. (PNo_bd_In) We take the following as an axiom:
∀L R : set(setprop)prop, PNoLt_pwise L R∀alpha, ordinal alphaPNo_lenbdd alpha LPNo_lenbdd alpha RPNo_bd L R ordsucc alpha
Beginning of Section TaggedSets
Let tag : setsetλalpha ⇒ SetAdjoin alpha {1}
Notation. We use ' as a postfix operator with priority 100 corresponding to applying term tag.
Axiom. (not_TransSet_Sing1) We take the following as an axiom:
Axiom. (not_ordinal_Sing1) We take the following as an axiom:
Axiom. (tagged_not_ordinal) We take the following as an axiom:
∀y, ¬ ordinal (y ')
Axiom. (tagged_notin_ordinal) We take the following as an axiom:
∀alpha y, ordinal alpha(y ') alpha
Axiom. (tagged_eqE_Subq) We take the following as an axiom:
∀alpha beta, ordinal alphaalpha ' = beta 'alpha beta
Axiom. (tagged_eqE_eq) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaalpha ' = beta 'alpha = beta
Axiom. (tagged_ReplE) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betabeta ' {gamma '|gammaalpha}beta alpha
Axiom. (ordinal_notin_tagged_Repl) We take the following as an axiom:
∀alpha Y, ordinal alphaalpha {y '|yY}
Definition. We define SNoElts_ to be λalpha ⇒ alpha {beta '|betaalpha} of type setset.
Axiom. (SNoElts_mon) We take the following as an axiom:
∀alpha beta, alpha betaSNoElts_ alpha SNoElts_ beta
Definition. We define SNo_ to be λalpha x ⇒ x SNoElts_ alpha betaalpha, exactly1of2 (beta ' x) (beta x) of type setsetprop.
Definition. We define PSNo to be λalpha p ⇒ {betaalpha|p beta} {beta '|betaalpha, ¬ p beta} of type set(setprop)set.
Axiom. (PNoEq_PSNo) We take the following as an axiom:
∀alpha, ordinal alpha∀p : setprop, PNoEq_ alpha (λbeta ⇒ beta PSNo alpha p) p
Axiom. (SNo_PSNo) We take the following as an axiom:
∀alpha, ordinal alpha∀p : setprop, SNo_ alpha (PSNo alpha p)
Axiom. (SNo_PSNo_eta_) We take the following as an axiom:
∀alpha x, ordinal alphaSNo_ alpha xx = PSNo alpha (λbeta ⇒ beta x)
Primitive. The name SNo is a term of type setprop.
Axiom. (SNo_SNo) We take the following as an axiom:
∀alpha, ordinal alpha∀z, SNo_ alpha zSNo z
Primitive. The name SNoLev is a term of type setset.
Axiom. (SNoLev_uniq_Subq) We take the following as an axiom:
∀x alpha beta, ordinal alphaordinal betaSNo_ alpha xSNo_ beta xalpha beta
Axiom. (SNoLev_uniq) We take the following as an axiom:
∀x alpha beta, ordinal alphaordinal betaSNo_ alpha xSNo_ beta xalpha = beta
Axiom. (SNoLev_prop) We take the following as an axiom:
∀x, SNo xordinal (SNoLev x) SNo_ (SNoLev x) x
Axiom. (SNoLev_ordinal) We take the following as an axiom:
∀x, SNo xordinal (SNoLev x)
Axiom. (SNoLev_) We take the following as an axiom:
∀x, SNo xSNo_ (SNoLev x) x
Axiom. (SNo_PSNo_eta) We take the following as an axiom:
∀x, SNo xx = PSNo (SNoLev x) (λbeta ⇒ beta x)
Axiom. (SNoLev_PSNo) We take the following as an axiom:
∀alpha, ordinal alpha∀p : setprop, SNoLev (PSNo alpha p) = alpha
Axiom. (SNo_Subq) We take the following as an axiom:
∀x y, SNo xSNo ySNoLev x SNoLev y(alphaSNoLev x, alpha x alpha y)x y
Definition. We define SNoEq_ to be λalpha x y ⇒ PNoEq_ alpha (λbeta ⇒ beta x) (λbeta ⇒ beta y) of type setsetsetprop.
Axiom. (SNoEq_I) We take the following as an axiom:
∀alpha x y, (betaalpha, beta x beta y)SNoEq_ alpha x y
Axiom. (SNo_eq) We take the following as an axiom:
∀x y, SNo xSNo ySNoLev x = SNoLev ySNoEq_ (SNoLev x) x yx = y
End of Section TaggedSets
Definition. We define SNoLt to be λx y ⇒ PNoLt (SNoLev x) (λbeta ⇒ beta x) (SNoLev y) (λbeta ⇒ beta y) of type setsetprop.
Notation. We use < as an infix operator with priority 490 and no associativity corresponding to applying term SNoLt.
Definition. We define SNoLe to be λx y ⇒ PNoLe (SNoLev x) (λbeta ⇒ beta x) (SNoLev y) (λbeta ⇒ beta y) of type setsetprop.
Notation. We use as an infix operator with priority 490 and no associativity corresponding to applying term SNoLe.
Axiom. (SNoLtLe) We take the following as an axiom:
∀x y, x < yx y
Axiom. (SNoLeE) We take the following as an axiom:
∀x y, SNo xSNo yx yx < y x = y
Axiom. (SNoEq_sym_) We take the following as an axiom:
∀alpha x y, SNoEq_ alpha x ySNoEq_ alpha y x
Axiom. (SNoEq_tra_) We take the following as an axiom:
∀alpha x y z, SNoEq_ alpha x ySNoEq_ alpha y zSNoEq_ alpha x z
Axiom. (SNoLtE) We take the following as an axiom:
∀x y, SNo xSNo yx < y∀p : prop, (∀z, SNo zSNoLev z SNoLev x SNoLev ySNoEq_ (SNoLev z) z xSNoEq_ (SNoLev z) z yx < zz < ySNoLev z xSNoLev z yp)(SNoLev x SNoLev ySNoEq_ (SNoLev x) x ySNoLev x yp)(SNoLev y SNoLev xSNoEq_ (SNoLev y) x ySNoLev y xp)p
Axiom. (SNoLtI2) We take the following as an axiom:
∀x y, SNoLev x SNoLev ySNoEq_ (SNoLev x) x ySNoLev x yx < y
Axiom. (SNoLtI3) We take the following as an axiom:
∀x y, SNoLev y SNoLev xSNoEq_ (SNoLev y) x ySNoLev y xx < y
Axiom. (SNoLt_irref) We take the following as an axiom:
∀x, ¬ SNoLt x x
Axiom. (SNoLt_trichotomy_or) We take the following as an axiom:
∀x y, SNo xSNo yx < y x = y y < x
Axiom. (SNoLt_trichotomy_or_impred) We take the following as an axiom:
∀x y, SNo xSNo y∀p : prop, (x < yp)(x = yp)(y < xp)p
Axiom. (SNoLt_tra) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx < yy < zx < z
Axiom. (SNoLe_ref) We take the following as an axiom:
∀x, SNoLe x x
Axiom. (SNoLe_antisym) We take the following as an axiom:
∀x y, SNo xSNo yx yy xx = y
Axiom. (SNoLtLe_tra) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx < yy zx < z
Axiom. (SNoLeLt_tra) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx yy < zx < z
Axiom. (SNoLe_tra) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx yy zx z
Axiom. (SNoLtLe_or) We take the following as an axiom:
∀x y, SNo xSNo yx < y y x
Axiom. (SNoLt_PSNo_PNoLt) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, ordinal alphaordinal betaPSNo alpha p < PSNo beta qPNoLt alpha p beta q
Axiom. (PNoLt_SNoLt_PSNo) We take the following as an axiom:
∀alpha beta, ∀p q : setprop, ordinal alphaordinal betaPNoLt alpha p beta qPSNo alpha p < PSNo beta q
Definition. We define SNoCut to be λL R ⇒ PSNo (PNo_bd (λalpha p ⇒ ordinal alpha PSNo alpha p L) (λalpha p ⇒ ordinal alpha PSNo alpha p R)) (PNo_pred (λalpha p ⇒ ordinal alpha PSNo alpha p L) (λalpha p ⇒ ordinal alpha PSNo alpha p R)) of type setsetset.
Definition. We define SNoCutP to be λL R ⇒ (xL, SNo x) (yR, SNo y) (xL, yR, x < y) of type setsetprop.
Axiom. (SNoCutP_SNoCut) We take the following as an axiom:
∀L R, SNoCutP L RSNo (SNoCut L R) SNoLev (SNoCut L R) ordsucc ((xLordsucc (SNoLev x)) (yRordsucc (SNoLev y))) (xL, x < SNoCut L R) (yR, SNoCut L R < y) (∀z, SNo z(xL, x < z)(yR, z < y)SNoLev (SNoCut L R) SNoLev z SNoEq_ (SNoLev (SNoCut L R)) (SNoCut L R) z)
Axiom. (SNoCutP_SNoCut_impred) We take the following as an axiom:
∀L R, SNoCutP L R∀p : prop, (SNo (SNoCut L R)SNoLev (SNoCut L R) ordsucc ((xLordsucc (SNoLev x)) (yRordsucc (SNoLev y)))(xL, x < SNoCut L R)(yR, SNoCut L R < y)(∀z, SNo z(xL, x < z)(yR, z < y)SNoLev (SNoCut L R) SNoLev z SNoEq_ (SNoLev (SNoCut L R)) (SNoCut L R) z)p)p
Axiom. (SNoCutP_L_0) We take the following as an axiom:
∀L, (xL, SNo x)SNoCutP L 0
Axiom. (SNoCutP_0_R) We take the following as an axiom:
∀R, (xR, SNo x)SNoCutP 0 R
Axiom. (SNoCutP_0_0) We take the following as an axiom:
SNoCutP 0 0
Definition. We define SNoS_ to be λalpha ⇒ {x𝒫 (SNoElts_ alpha)|betaalpha, SNo_ beta x} of type setset.
Axiom. (SNoS_E) We take the following as an axiom:
∀alpha, ordinal alphaxSNoS_ alpha, betaalpha, SNo_ beta x
Beginning of Section TaggedSets2
Let tag : setsetλalpha ⇒ SetAdjoin alpha {1}
Notation. We use ' as a postfix operator with priority 100 corresponding to applying term tag.
Axiom. (SNoS_I) We take the following as an axiom:
∀alpha, ordinal alpha∀x, betaalpha, SNo_ beta xx SNoS_ alpha
Axiom. (SNoS_I2) We take the following as an axiom:
∀x y, SNo xSNo ySNoLev x SNoLev yx SNoS_ (SNoLev y)
Axiom. (SNoS_Subq) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaalpha betaSNoS_ alpha SNoS_ beta
Axiom. (SNoLev_uniq2) We take the following as an axiom:
∀alpha, ordinal alpha∀x, SNo_ alpha xSNoLev x = alpha
Axiom. (SNoS_E2) We take the following as an axiom:
∀alpha, ordinal alphaxSNoS_ alpha, ∀p : prop, (SNoLev x alphaordinal (SNoLev x)SNo xSNo_ (SNoLev x) xp)p
Axiom. (SNoS_In_neq) We take the following as an axiom:
∀w, SNo wxSNoS_ (SNoLev w), x w
Axiom. (SNoS_SNoLev) We take the following as an axiom:
∀z, SNo zz SNoS_ (ordsucc (SNoLev z))
Definition. We define SNoL to be λz ⇒ {xSNoS_ (SNoLev z)|x < z} of type setset.
Definition. We define SNoR to be λz ⇒ {ySNoS_ (SNoLev z)|z < y} of type setset.
Axiom. (SNoCutP_SNoL_SNoR) We take the following as an axiom:
∀z, SNo zSNoCutP (SNoL z) (SNoR z)
Axiom. (SNoL_E) We take the following as an axiom:
∀x, SNo xwSNoL x, ∀p : prop, (SNo wSNoLev w SNoLev xw < xp)p
Axiom. (SNoR_E) We take the following as an axiom:
∀x, SNo xzSNoR x, ∀p : prop, (SNo zSNoLev z SNoLev xx < zp)p
Axiom. (SNoL_SNoS_) We take the following as an axiom:
∀z, SNoL z SNoS_ (SNoLev z)
Axiom. (SNoR_SNoS_) We take the following as an axiom:
∀z, SNoR z SNoS_ (SNoLev z)
Axiom. (SNoL_SNoS) We take the following as an axiom:
∀x, SNo xwSNoL x, w SNoS_ (SNoLev x)
Axiom. (SNoR_SNoS) We take the following as an axiom:
∀x, SNo xzSNoR x, z SNoS_ (SNoLev x)
Axiom. (SNoL_I) We take the following as an axiom:
∀z, SNo z∀x, SNo xSNoLev x SNoLev zx < zx SNoL z
Axiom. (SNoR_I) We take the following as an axiom:
∀z, SNo z∀y, SNo ySNoLev y SNoLev zz < yy SNoR z
Axiom. (SNo_eta) We take the following as an axiom:
∀z, SNo zz = SNoCut (SNoL z) (SNoR z)
Axiom. (SNoCutP_SNo_SNoCut) We take the following as an axiom:
∀L R, SNoCutP L RSNo (SNoCut L R)
Axiom. (SNoCutP_SNoCut_L) We take the following as an axiom:
∀L R, SNoCutP L RxL, x < SNoCut L R
Axiom. (SNoCutP_SNoCut_R) We take the following as an axiom:
∀L R, SNoCutP L RyR, SNoCut L R < y
Axiom. (SNoCutP_SNoCut_fst) We take the following as an axiom:
∀L R, SNoCutP L R∀z, SNo z(xL, x < z)(yR, z < y)SNoLev (SNoCut L R) SNoLev z SNoEq_ (SNoLev (SNoCut L R)) (SNoCut L R) z
Axiom. (SNoCut_Le) We take the following as an axiom:
∀L1 R1 L2 R2, SNoCutP L1 R1SNoCutP L2 R2(wL1, w < SNoCut L2 R2)(zR2, SNoCut L1 R1 < z)SNoCut L1 R1 SNoCut L2 R2
Axiom. (SNoCut_ext) We take the following as an axiom:
∀L1 R1 L2 R2, SNoCutP L1 R1SNoCutP L2 R2(wL1, w < SNoCut L2 R2)(zR1, SNoCut L2 R2 < z)(wL2, w < SNoCut L1 R1)(zR2, SNoCut L1 R1 < z)SNoCut L1 R1 = SNoCut L2 R2
Axiom. (SNoLt_SNoL_or_SNoR_impred) We take the following as an axiom:
∀x y, SNo xSNo yx < y∀p : prop, (zSNoL y, z SNoR xp)(x SNoL yp)(y SNoR xp)p
Axiom. (SNoL_or_SNoR_impred) We take the following as an axiom:
∀x y, SNo xSNo y∀p : prop, (x = yp)(zSNoL y, z SNoR xp)(x SNoL yp)(y SNoR xp)(zSNoR y, z SNoL xp)(x SNoR yp)(y SNoL xp)p
Axiom. (ordinal_SNo_) We take the following as an axiom:
∀alpha, ordinal alphaSNo_ alpha alpha
Axiom. (ordinal_SNo) We take the following as an axiom:
∀alpha, ordinal alphaSNo alpha
Axiom. (ordinal_SNoLev) We take the following as an axiom:
∀alpha, ordinal alphaSNoLev alpha = alpha
Axiom. (ordinal_SNoLev_max) We take the following as an axiom:
∀alpha, ordinal alpha∀z, SNo zSNoLev z alphaz < alpha
Axiom. (ordinal_SNoL) We take the following as an axiom:
∀alpha, ordinal alphaSNoL alpha = SNoS_ alpha
Axiom. (ordinal_SNoR) We take the following as an axiom:
∀alpha, ordinal alphaSNoR alpha = Empty
Axiom. (nat_p_SNo) We take the following as an axiom:
∀n, nat_p nSNo n
Axiom. (omega_SNo) We take the following as an axiom:
Axiom. (omega_SNoS_omega) We take the following as an axiom:
Axiom. (ordinal_In_SNoLt) We take the following as an axiom:
∀alpha, ordinal alphabetaalpha, beta < alpha
Axiom. (ordinal_SNoLev_max_2) We take the following as an axiom:
∀alpha, ordinal alpha∀z, SNo zSNoLev z ordsucc alphaz alpha
Axiom. (ordinal_Subq_SNoLe) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaalpha betaalpha beta
Axiom. (ordinal_SNoLt_In) We take the following as an axiom:
∀alpha beta, ordinal alphaordinal betaalpha < betaalpha beta
Axiom. (omega_nonneg) We take the following as an axiom:
Axiom. (SNo_0) We take the following as an axiom:
Axiom. (SNo_1) We take the following as an axiom:
Axiom. (SNo_2) We take the following as an axiom:
Axiom. (SNoLev_0) We take the following as an axiom:
Axiom. (SNoCut_0_0) We take the following as an axiom:
Axiom. (SNoL_0) We take the following as an axiom:
Axiom. (SNoR_0) We take the following as an axiom:
Axiom. (SNoL_1) We take the following as an axiom:
Axiom. (SNoR_1) We take the following as an axiom:
Axiom. (SNo_max_SNoLev) We take the following as an axiom:
∀x, SNo x(ySNoS_ (SNoLev x), y < x)SNoLev x = x
Axiom. (SNo_max_ordinal) We take the following as an axiom:
∀x, SNo x(ySNoS_ (SNoLev x), y < x)ordinal x
Definition. We define SNo_extend0 to be λx ⇒ PSNo (ordsucc (SNoLev x)) (λdelta ⇒ delta x delta SNoLev x) of type setset.
Definition. We define SNo_extend1 to be λx ⇒ PSNo (ordsucc (SNoLev x)) (λdelta ⇒ delta x delta = SNoLev x) of type setset.
Axiom. (SNo_extend0_SNo_) We take the following as an axiom:
∀x, SNo xSNo_ (ordsucc (SNoLev x)) (SNo_extend0 x)
Axiom. (SNo_extend1_SNo_) We take the following as an axiom:
∀x, SNo xSNo_ (ordsucc (SNoLev x)) (SNo_extend1 x)
Axiom. (SNo_extend0_SNo) We take the following as an axiom:
∀x, SNo xSNo (SNo_extend0 x)
Axiom. (SNo_extend1_SNo) We take the following as an axiom:
∀x, SNo xSNo (SNo_extend1 x)
Axiom. (SNo_extend0_SNoLev) We take the following as an axiom:
∀x, SNo xSNoLev (SNo_extend0 x) = ordsucc (SNoLev x)
Axiom. (SNo_extend1_SNoLev) We take the following as an axiom:
∀x, SNo xSNoLev (SNo_extend1 x) = ordsucc (SNoLev x)
Axiom. (SNo_extend0_nIn) We take the following as an axiom:
∀x, SNo xSNoLev x SNo_extend0 x
Axiom. (SNo_extend1_In) We take the following as an axiom:
∀x, SNo xSNoLev x SNo_extend1 x
Axiom. (SNo_extend0_SNoEq) We take the following as an axiom:
∀x, SNo xSNoEq_ (SNoLev x) (SNo_extend0 x) x
Axiom. (SNo_extend1_SNoEq) We take the following as an axiom:
∀x, SNo xSNoEq_ (SNoLev x) (SNo_extend1 x) x
Axiom. (SNoLev_0_eq_0) We take the following as an axiom:
∀x, SNo xSNoLev x = 0x = 0
Definition. We define eps_ to be λn ⇒ {0} {(ordsucc m) '|mn} of type setset.
Axiom. (eps_ordinal_In_eq_0) We take the following as an axiom:
∀n alpha, ordinal alphaalpha eps_ nalpha = 0
Axiom. (eps_0_1) We take the following as an axiom:
Axiom. (SNo__eps_) We take the following as an axiom:
Axiom. (SNo_eps_) We take the following as an axiom:
Axiom. (SNo_eps_1) We take the following as an axiom:
Axiom. (SNoLev_eps_) We take the following as an axiom:
Axiom. (SNo_eps_SNoS_omega) We take the following as an axiom:
Axiom. (SNo_eps_decr) We take the following as an axiom:
Axiom. (SNo_eps_pos) We take the following as an axiom:
Axiom. (SNo_pos_eps_Lt) We take the following as an axiom:
∀n, nat_p nxSNoS_ (ordsucc n), 0 < xeps_ n < x
Axiom. (SNo_pos_eps_Le) We take the following as an axiom:
∀n, nat_p nxSNoS_ (ordsucc (ordsucc n)), 0 < xeps_ n x
Axiom. (eps_SNo_eq) We take the following as an axiom:
∀n, nat_p nxSNoS_ (ordsucc n), 0 < xSNoEq_ (SNoLev x) (eps_ n) xmn, x = eps_ m
Axiom. (eps_SNoCutP) We take the following as an axiom:
Axiom. (eps_SNoCut) We take the following as an axiom:
End of Section TaggedSets2
Axiom. (SNo_etaE) We take the following as an axiom:
∀z, SNo z∀p : prop, (∀L R, SNoCutP L R(xL, SNoLev x SNoLev z)(yR, SNoLev y SNoLev z)z = SNoCut L Rp)p
Axiom. (SNo_ind) We take the following as an axiom:
∀P : setprop, (∀L R, SNoCutP L R(xL, P x)(yR, P y)P (SNoCut L R))∀z, SNo zP z
Beginning of Section SurrealRecI
Variable F : set(setset)set
Let default : setEps_i (λ_ ⇒ True)
Let G : set(setsetset)setsetλalpha g ⇒ If_ii (ordinal alpha) (λz : setif z SNoS_ (ordsucc alpha) then F z (λw ⇒ g (SNoLev w) w) else default) (λz : setdefault)
Primitive. The name SNo_rec_i is a term of type setset.
Hypothesis Fr : ∀z, SNo z∀g h : setset, (wSNoS_ (SNoLev z), g w = h w)F z g = F z h
Axiom. (SNo_rec_i_eq) We take the following as an axiom:
∀z, SNo zSNo_rec_i z = F z SNo_rec_i
End of Section SurrealRecI
Beginning of Section SurrealRecII
Variable F : set(set(setset))(setset)
Let default : (setset)Descr_ii (λ_ ⇒ True)
Let G : set(setset(setset))set(setset)λalpha g ⇒ If_iii (ordinal alpha) (λz : setIf_ii (z SNoS_ (ordsucc alpha)) (F z (λw ⇒ g (SNoLev w) w)) default) (λz : setdefault)
Primitive. The name SNo_rec_ii is a term of type set(setset).
Hypothesis Fr : ∀z, SNo z∀g h : set(setset), (wSNoS_ (SNoLev z), g w = h w)F z g = F z h
Axiom. (SNo_rec_ii_eq) We take the following as an axiom:
∀z, SNo zSNo_rec_ii z = F z SNo_rec_ii
End of Section SurrealRecII
Beginning of Section SurrealRec2
Variable F : setset(setsetset)set
Let G : set(setsetset)set(setset)setλw f z g ⇒ F w z (λx y ⇒ if x = w then g y else f x y)
Let H : set(setsetset)setsetλw f z ⇒ if SNo z then SNo_rec_i (G w f) z else Empty
Primitive. The name SNo_rec2 is a term of type setsetset.
Hypothesis Fr : ∀w, SNo w∀z, SNo z∀g h : setsetset, (xSNoS_ (SNoLev w), ∀y, SNo yg x y = h x y)(ySNoS_ (SNoLev z), g w y = h w y)F w z g = F w z h
Axiom. (SNo_rec2_G_prop) We take the following as an axiom:
∀w, SNo w∀f k : setsetset, (xSNoS_ (SNoLev w), f x = k x)∀z, SNo z∀g h : setset, (uSNoS_ (SNoLev z), g u = h u)G w f z g = G w k z h
Axiom. (SNo_rec2_eq_1) We take the following as an axiom:
∀w, SNo w∀f : setsetset, ∀z, SNo zSNo_rec_i (G w f) z = G w f z (SNo_rec_i (G w f))
Axiom. (SNo_rec2_eq) We take the following as an axiom:
∀w, SNo w∀z, SNo zSNo_rec2 w z = F w z SNo_rec2
End of Section SurrealRec2
Axiom. (SNo_ordinal_ind) We take the following as an axiom:
∀P : setprop, (∀alpha, ordinal alphaxSNoS_ alpha, P x)(∀x, SNo xP x)
Axiom. (SNo_ordinal_ind2) We take the following as an axiom:
∀P : setsetprop, (∀alpha, ordinal alpha∀beta, ordinal betaxSNoS_ alpha, ySNoS_ beta, P x y)(∀x y, SNo xSNo yP x y)
Axiom. (SNo_ordinal_ind3) We take the following as an axiom:
∀P : setsetsetprop, (∀alpha, ordinal alpha∀beta, ordinal beta∀gamma, ordinal gammaxSNoS_ alpha, ySNoS_ beta, zSNoS_ gamma, P x y z)(∀x y z, SNo xSNo ySNo zP x y z)
Axiom. (SNoLev_ind) We take the following as an axiom:
∀P : setprop, (∀x, SNo x(wSNoS_ (SNoLev x), P w)P x)(∀x, SNo xP x)
Axiom. (SNoLev_ind2) We take the following as an axiom:
∀P : setsetprop, (∀x y, SNo xSNo y(wSNoS_ (SNoLev x), P w y)(zSNoS_ (SNoLev y), P x z)(wSNoS_ (SNoLev x), zSNoS_ (SNoLev y), P w z)P x y)∀x y, SNo xSNo yP x y
Axiom. (SNoLev_ind3) We take the following as an axiom:
∀P : setsetsetprop, (∀x y z, SNo xSNo ySNo z(uSNoS_ (SNoLev x), P u y z)(vSNoS_ (SNoLev y), P x v z)(wSNoS_ (SNoLev z), P x y w)(uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), P u v z)(uSNoS_ (SNoLev x), wSNoS_ (SNoLev z), P u y w)(vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), P x v w)(uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), P u v w)P x y z)∀x y z, SNo xSNo ySNo zP x y z
Axiom. (SNo_omega) We take the following as an axiom:
SNo ω
Axiom. (SNoLt_0_1) We take the following as an axiom:
0 < 1
Axiom. (SNoLt_0_2) We take the following as an axiom:
0 < 2
Axiom. (SNoLt_1_2) We take the following as an axiom:
1 < 2
Axiom. (restr_SNo_) We take the following as an axiom:
∀x, SNo xalphaSNoLev x, SNo_ alpha (x SNoElts_ alpha)
Axiom. (restr_SNo) We take the following as an axiom:
∀x, SNo xalphaSNoLev x, SNo (x SNoElts_ alpha)
Axiom. (restr_SNoLev) We take the following as an axiom:
∀x, SNo xalphaSNoLev x, SNoLev (x SNoElts_ alpha) = alpha
Axiom. (restr_SNoEq) We take the following as an axiom:
∀x, SNo xalphaSNoLev x, SNoEq_ alpha (x SNoElts_ alpha) x
Axiom. (SNo_extend0_restr_eq) We take the following as an axiom:
∀x, SNo xx = SNo_extend0 x SNoElts_ (SNoLev x)
Axiom. (SNo_extend1_restr_eq) We take the following as an axiom:
∀x, SNo xx = SNo_extend1 x SNoElts_ (SNoLev x)
Beginning of Section SurrealMinus
Primitive. The name minus_SNo is a term of type setset.
Notation. We use - as a prefix operator with priority 358 corresponding to applying term minus_SNo.
Notation. We use as an infix operator with priority 490 and no associativity corresponding to applying term SNoLe.
Axiom. (minus_SNo_eq) We take the following as an axiom:
∀x, SNo x- x = SNoCut {- z|zSNoR x} {- w|wSNoL x}
Axiom. (minus_SNo_prop1) We take the following as an axiom:
∀x, SNo xSNo (- x) (uSNoL x, - x < - u) (uSNoR x, - u < - x) SNoCutP {- z|zSNoR x} {- w|wSNoL x}
Axiom. (SNo_minus_SNo) We take the following as an axiom:
∀x, SNo xSNo (- x)
Axiom. (minus_SNo_Lt_contra) We take the following as an axiom:
∀x y, SNo xSNo yx < y- y < - x
Axiom. (minus_SNo_Le_contra) We take the following as an axiom:
∀x y, SNo xSNo yx y- y - x
Axiom. (minus_SNo_SNoCutP) We take the following as an axiom:
∀x, SNo xSNoCutP {- z|zSNoR x} {- w|wSNoL x}
Axiom. (minus_SNo_SNoCutP_gen) We take the following as an axiom:
∀L R, SNoCutP L RSNoCutP {- z|zR} {- w|wL}
Axiom. (minus_SNo_Lev_lem1) We take the following as an axiom:
∀alpha, ordinal alphaxSNoS_ alpha, SNoLev (- x) SNoLev x
Axiom. (minus_SNo_Lev_lem2) We take the following as an axiom:
∀x, SNo xSNoLev (- x) SNoLev x
Axiom. (minus_SNo_invol) We take the following as an axiom:
∀x, SNo x- - x = x
Axiom. (minus_SNo_Lev) We take the following as an axiom:
∀x, SNo xSNoLev (- x) = SNoLev x
Axiom. (minus_SNo_SNo_) We take the following as an axiom:
∀alpha, ordinal alpha∀x, SNo_ alpha xSNo_ alpha (- x)
Axiom. (minus_SNo_SNoS_) We take the following as an axiom:
∀alpha, ordinal alpha∀x, x SNoS_ alpha- x SNoS_ alpha
Axiom. (minus_SNoCut_eq_lem) We take the following as an axiom:
∀v, SNo v∀L R, SNoCutP L Rv = SNoCut L R- v = SNoCut {- z|zR} {- w|wL}
Axiom. (minus_SNoCut_eq) We take the following as an axiom:
∀L R, SNoCutP L R- SNoCut L R = SNoCut {- z|zR} {- w|wL}
Axiom. (minus_SNo_Lt_contra1) We take the following as an axiom:
∀x y, SNo xSNo y- x < y- y < x
Axiom. (minus_SNo_Lt_contra2) We take the following as an axiom:
∀x y, SNo xSNo yx < - yy < - x
Axiom. (mordinal_SNoLev_min_2) We take the following as an axiom:
∀alpha, ordinal alpha∀z, SNo zSNoLev z ordsucc alpha- alpha z
Axiom. (minus_SNo_SNoS_omega) We take the following as an axiom:
Axiom. (SNoL_minus_SNoR) We take the following as an axiom:
∀x, SNo xSNoL (- x) = {- w|wSNoR x}
End of Section SurrealMinus
Beginning of Section SurrealAdd
Notation. We use - as a prefix operator with priority 358 corresponding to applying term minus_SNo.
Primitive. The name add_SNo is a term of type setsetset.
Notation. We use + as an infix operator with priority 360 and which associates to the right corresponding to applying term add_SNo.
Axiom. (add_SNo_eq) We take the following as an axiom:
∀x, SNo x∀y, SNo yx + y = SNoCut ({w + y|wSNoL x} {x + w|wSNoL y}) ({z + y|zSNoR x} {x + z|zSNoR y})
Axiom. (add_SNo_prop1) We take the following as an axiom:
∀x y, SNo xSNo ySNo (x + y) (uSNoL x, u + y < x + y) (uSNoR x, x + y < u + y) (uSNoL y, x + u < x + y) (uSNoR y, x + y < x + u) SNoCutP ({w + y|wSNoL x} {x + w|wSNoL y}) ({z + y|zSNoR x} {x + z|zSNoR y})
Axiom. (SNo_add_SNo) We take the following as an axiom:
∀x y, SNo xSNo ySNo (x + y)
Axiom. (SNo_add_SNo_3) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zSNo (x + y + z)
Axiom. (SNo_add_SNo_3c) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zSNo (x + y + - z)
Axiom. (SNo_add_SNo_4) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wSNo (x + y + z + w)
Axiom. (add_SNo_Lt1) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx < zx + y < z + y
Axiom. (add_SNo_Le1) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx zx + y z + y
Axiom. (add_SNo_Lt2) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zy < zx + y < x + z
Axiom. (add_SNo_Le2) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zy zx + y x + z
Axiom. (add_SNo_Lt3a) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wx < zy wx + y < z + w
Axiom. (add_SNo_Lt3b) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wx zy < wx + y < z + w
Axiom. (add_SNo_Lt3) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wx < zy < wx + y < z + w
Axiom. (add_SNo_Le3) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wx zy wx + y z + w
Axiom. (add_SNo_SNoCutP) We take the following as an axiom:
∀x y, SNo xSNo ySNoCutP ({w + y|wSNoL x} {x + w|wSNoL y}) ({z + y|zSNoR x} {x + z|zSNoR y})
Axiom. (add_SNo_com) We take the following as an axiom:
∀x y, SNo xSNo yx + y = y + x
Axiom. (add_SNo_0L) We take the following as an axiom:
∀x, SNo x0 + x = x
Axiom. (add_SNo_0R) We take the following as an axiom:
∀x, SNo xx + 0 = x
Axiom. (add_SNo_minus_SNo_linv) We take the following as an axiom:
∀x, SNo x- x + x = 0
Axiom. (add_SNo_minus_SNo_rinv) We take the following as an axiom:
∀x, SNo xx + - x = 0
Axiom. (add_SNo_ordinal_SNoCutP) We take the following as an axiom:
∀alpha, ordinal alpha∀beta, ordinal betaSNoCutP ({x + beta|xSNoS_ alpha} {alpha + x|xSNoS_ beta}) Empty
Axiom. (add_SNo_ordinal_eq) We take the following as an axiom:
∀alpha, ordinal alpha∀beta, ordinal betaalpha + beta = SNoCut ({x + beta|xSNoS_ alpha} {alpha + x|xSNoS_ beta}) Empty
Axiom. (add_SNo_ordinal_ordinal) We take the following as an axiom:
∀alpha, ordinal alpha∀beta, ordinal betaordinal (alpha + beta)
Axiom. (add_SNo_ordinal_SL) We take the following as an axiom:
∀alpha, ordinal alpha∀beta, ordinal betaordsucc alpha + beta = ordsucc (alpha + beta)
Axiom. (add_SNo_ordinal_SR) We take the following as an axiom:
∀alpha, ordinal alpha∀beta, ordinal betaalpha + ordsucc beta = ordsucc (alpha + beta)
Axiom. (add_SNo_ordinal_InL) We take the following as an axiom:
∀alpha, ordinal alpha∀beta, ordinal betagammaalpha, gamma + beta alpha + beta
Axiom. (add_SNo_ordinal_InR) We take the following as an axiom:
∀alpha, ordinal alpha∀beta, ordinal betagammabeta, alpha + gamma alpha + beta
Axiom. (add_nat_add_SNo) We take the following as an axiom:
Axiom. (add_SNo_In_omega) We take the following as an axiom:
Axiom. (add_SNo_1_1_2) We take the following as an axiom:
1 + 1 = 2
Axiom. (add_SNo_SNoL_interpolate) We take the following as an axiom:
∀x y, SNo xSNo yuSNoL (x + y), (vSNoL x, u v + y) (vSNoL y, u x + v)
Axiom. (add_SNo_SNoR_interpolate) We take the following as an axiom:
∀x y, SNo xSNo yuSNoR (x + y), (vSNoR x, v + y u) (vSNoR y, x + v u)
Axiom. (add_SNo_assoc) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + (y + z) = (x + y) + z
Axiom. (add_SNo_cancel_L) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + y = x + zy = z
Axiom. (minus_SNo_0) We take the following as an axiom:
- 0 = 0
Axiom. (minus_add_SNo_distr) We take the following as an axiom:
∀x y, SNo xSNo y- (x + y) = (- x) + (- y)
Axiom. (minus_add_SNo_distr_3) We take the following as an axiom:
∀x y z, SNo xSNo ySNo z- (x + y + z) = - x + - y + - z
Axiom. (add_SNo_Lev_bd) We take the following as an axiom:
∀x y, SNo xSNo ySNoLev (x + y) SNoLev x + SNoLev y
Axiom. (add_SNo_SNoS_omega) We take the following as an axiom:
Axiom. (add_SNo_minus_R2) We take the following as an axiom:
∀x y, SNo xSNo y(x + y) + - y = x
Axiom. (add_SNo_minus_R2') We take the following as an axiom:
∀x y, SNo xSNo y(x + - y) + y = x
Axiom. (add_SNo_minus_L2) We take the following as an axiom:
∀x y, SNo xSNo y- x + (x + y) = y
Axiom. (add_SNo_minus_L2') We take the following as an axiom:
∀x y, SNo xSNo yx + (- x + y) = y
Axiom. (add_SNo_Lt1_cancel) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + y < z + yx < z
Axiom. (add_SNo_Lt2_cancel) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + y < x + zy < z
Axiom. (add_SNo_assoc_4) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wx + y + z + w = (x + y + z) + w
Axiom. (add_SNo_com_3_0_1) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + y + z = y + x + z
Axiom. (add_SNo_com_3b_1_2) We take the following as an axiom:
∀x y z, SNo xSNo ySNo z(x + y) + z = (x + z) + y
Axiom. (add_SNo_com_4_inner_mid) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w(x + y) + (z + w) = (x + z) + (y + w)
Axiom. (add_SNo_rotate_3_1) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + y + z = z + x + y
Axiom. (add_SNo_rotate_4_1) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wx + y + z + w = w + x + y + z
Axiom. (add_SNo_rotate_5_1) We take the following as an axiom:
∀x y z w v, SNo xSNo ySNo zSNo wSNo vx + y + z + w + v = v + x + y + z + w
Axiom. (add_SNo_rotate_5_2) We take the following as an axiom:
∀x y z w v, SNo xSNo ySNo zSNo wSNo vx + y + z + w + v = w + v + x + y + z
Axiom. (add_SNo_minus_SNo_prop3) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w(x + y + z) + (- z + w) = x + y + w
Axiom. (add_SNo_minus_SNo_prop4) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w(x + y + z) + (w + - z) = x + y + w
Axiom. (add_SNo_minus_SNo_prop5) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w(x + y + - z) + (z + w) = x + y + w
Axiom. (add_SNo_minus_Lt1) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + - y < zx < z + y
Axiom. (add_SNo_minus_Lt2) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zz < x + - yz + y < x
Axiom. (add_SNo_minus_Lt1b) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx < z + yx + - y < z
Axiom. (add_SNo_minus_Lt2b) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zz + y < xz < x + - y
Axiom. (add_SNo_minus_Lt1b3) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo wx + y < w + zx + y + - z < w
Axiom. (add_SNo_minus_Lt2b3) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo ww + z < x + yw < x + y + - z
Axiom. (add_SNo_minus_Lt_lem) We take the following as an axiom:
∀x y z u v w, SNo xSNo ySNo zSNo uSNo vSNo wx + y + w < u + v + zx + y + - z < u + v + - w
Axiom. (add_SNo_minus_Le2) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zz x + - yz + y x
Axiom. (add_SNo_minus_Le2b) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zz + y xz x + - y
Axiom. (add_SNo_Lt_subprop2) We take the following as an axiom:
∀x y z w u v, SNo xSNo ySNo zSNo wSNo uSNo vx + u < z + vy + v < w + ux + y < z + w
Axiom. (add_SNo_Lt_subprop3a) We take the following as an axiom:
∀x y z w u a, SNo xSNo ySNo zSNo wSNo uSNo ax + z < w + ay + a < ux + y + z < w + u
Axiom. (add_SNo_Lt_subprop3b) We take the following as an axiom:
∀x y w u v a, SNo xSNo ySNo wSNo uSNo vSNo ax + a < w + vy < a + ux + y < w + u + v
Axiom. (add_SNo_Lt_subprop3c) We take the following as an axiom:
∀x y z w u a b c, SNo xSNo ySNo zSNo wSNo uSNo aSNo bSNo cx + a < b + cy + c < ub + z < w + ax + y + z < w + u
Axiom. (add_SNo_Lt_subprop3d) We take the following as an axiom:
∀x y w u v a b c, SNo xSNo ySNo wSNo uSNo vSNo aSNo bSNo cx + a < b + vy < c + ub + c < w + ax + y < w + u + v
Axiom. (ordinal_ordsucc_SNo_eq) We take the following as an axiom:
∀alpha, ordinal alphaordsucc alpha = 1 + alpha
Axiom. (add_SNo_3a_2b) We take the following as an axiom:
∀x y z w u, SNo xSNo ySNo zSNo wSNo u(x + y + z) + (w + u) = (u + y + z) + (w + x)
Axiom. (add_SNo_1_ordsucc) We take the following as an axiom:
Axiom. (add_SNo_eps_Lt) We take the following as an axiom:
∀x, SNo xnω, x < x + eps_ n
Axiom. (add_SNo_eps_Lt') We take the following as an axiom:
∀x y, SNo xSNo ynω, x < yx < y + eps_ n
Axiom. (SNoLt_minus_pos) We take the following as an axiom:
∀x y, SNo xSNo yx < y0 < y + - x
Axiom. (add_SNo_omega_In_cases) We take the following as an axiom:
∀m, nω, ∀k, nat_p km n + km n m + - n k
Axiom. (add_SNo_Lt4) We take the following as an axiom:
∀x y z w u v, SNo xSNo ySNo zSNo wSNo uSNo vx < wy < uz < vx + y + z < w + u + v
Axiom. (add_SNo_3_3_3_Lt1) We take the following as an axiom:
∀x y z w u, SNo xSNo ySNo zSNo wSNo ux + y < z + wx + y + u < z + w + u
Axiom. (add_SNo_3_2_3_Lt1) We take the following as an axiom:
∀x y z w u, SNo xSNo ySNo zSNo wSNo uy + x < z + wx + u + y < z + w + u
Axiom. (add_SNo_minus_Lt12b3) We take the following as an axiom:
∀x y z w u v, SNo xSNo ySNo zSNo wSNo uSNo vx + y + v < w + u + zx + y + - z < w + u + - v
Axiom. (add_SNo_minus_Le12b3) We take the following as an axiom:
∀x y z w u v, SNo xSNo ySNo zSNo wSNo uSNo vx + y + v w + u + zx + y + - z w + u + - v
End of Section SurrealAdd
Beginning of Section SurrealMul
Notation. We use - as a prefix operator with priority 358 corresponding to applying term minus_SNo.
Notation. We use + as an infix operator with priority 360 and which associates to the right corresponding to applying term add_SNo.
Definition. We define mul_SNo to be SNo_rec2 (λx y m ⇒ SNoCut ({m (w 0) y + m x (w 1) + - m (w 0) (w 1)|wSNoL x SNoL y} {m (z 0) y + m x (z 1) + - m (z 0) (z 1)|zSNoR x SNoR y}) ({m (w 0) y + m x (w 1) + - m (w 0) (w 1)|wSNoL x SNoR y} {m (z 0) y + m x (z 1) + - m (z 0) (z 1)|zSNoR x SNoL y})) of type setsetset.
Notation. We use * as an infix operator with priority 355 and which associates to the right corresponding to applying term mul_SNo.
Theorem. (mul_SNo_eq) The following is provable:
∀x, SNo x∀y, SNo yx * y = SNoCut ({(w 0) * y + x * (w 1) + - (w 0) * (w 1)|wSNoL x SNoL y} {(z 0) * y + x * (z 1) + - (z 0) * (z 1)|zSNoR x SNoR y}) ({(w 0) * y + x * (w 1) + - (w 0) * (w 1)|wSNoL x SNoR y} {(z 0) * y + x * (z 1) + - (z 0) * (z 1)|zSNoR x SNoL y})
Proof:
Set F to be the term λx y m ⇒ SNoCut ({m (w 0) y + m x (w 1) + - m (w 0) (w 1)|wSNoL x SNoL y} {m (z 0) y + m x (z 1) + - m (z 0) (z 1)|zSNoR x SNoR y}) ({m (w 0) y + m x (w 1) + - m (w 0) (w 1)|wSNoL x SNoR y} {m (z 0) y + m x (z 1) + - m (z 0) (z 1)|zSNoR x SNoL y}) of type setset(setsetset)set.
We prove the intermediate claim L1: ∀x, SNo x∀y, SNo y∀g h : setsetset, (wSNoS_ (SNoLev x), ∀z, SNo zg w z = h w z)(zSNoS_ (SNoLev y), g x z = h x z)F x y g = F x y h.
Let x be given.
Assume Hx: SNo x.
Let y be given.
Assume Hy: SNo y.
Let g and h be given.
Assume Hgh1: wSNoS_ (SNoLev x), ∀z, SNo zg w z = h w z.
Assume Hgh2: zSNoS_ (SNoLev y), g x z = h x z.
We will prove F x y g = F x y h.
We prove the intermediate claim L1a: {g (w 0) y + g x (w 1) + - g (w 0) (w 1)|wSNoL x SNoL y} = {h (w 0) y + h x (w 1) + - h (w 0) (w 1)|wSNoL x SNoL y}.
Apply ReplEq_setprod_ext (SNoL x) (SNoL y) (λw0 w1 ⇒ g w0 y + g x w1 + - g w0 w1) (λw0 w1 ⇒ h w0 y + h x w1 + - h w0 w1) to the current goal.
We will prove w0SNoL x, w1SNoL y, g w0 y + g x w1 + - g w0 w1 = h w0 y + h x w1 + - h w0 w1.
Let w0 be given.
Assume Hw0: w0 SNoL x.
Let w1 be given.
Assume Hw1: w1 SNoL y.
We prove the intermediate claim Lw0: w0 SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoL_SNoS x Hx w0 Hw0.
We prove the intermediate claim Lw1: w1 SNoS_ (SNoLev y).
An exact proof term for the current goal is SNoL_SNoS y Hy w1 Hw1.
We prove the intermediate claim Lw1b: SNo w1.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim L1aa: g w0 y = h w0 y.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0.
An exact proof term for the current goal is Hy.
We prove the intermediate claim L1ab: g x w1 = h x w1.
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lw1.
We prove the intermediate claim L1ac: g w0 w1 = h w0 w1.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0.
An exact proof term for the current goal is Lw1b.
rewrite the current goal using L1aa (from left to right).
rewrite the current goal using L1ab (from left to right).
rewrite the current goal using L1ac (from left to right).
Use reflexivity.
We prove the intermediate claim L1b: {g (z 0) y + g x (z 1) + - g (z 0) (z 1)|zSNoR x SNoR y} = {h (z 0) y + h x (z 1) + - h (z 0) (z 1)|zSNoR x SNoR y}.
Apply ReplEq_setprod_ext (SNoR x) (SNoR y) (λz0 z1 ⇒ g z0 y + g x z1 + - g z0 z1) (λz0 z1 ⇒ h z0 y + h x z1 + - h z0 z1) to the current goal.
We will prove z0SNoR x, z1SNoR y, g z0 y + g x z1 + - g z0 z1 = h z0 y + h x z1 + - h z0 z1.
Let z0 be given.
Assume Hz0: z0 SNoR x.
Let z1 be given.
Assume Hz1: z1 SNoR y.
We prove the intermediate claim Lz0: z0 SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoR_SNoS x Hx z0 Hz0.
We prove the intermediate claim Lz1: z1 SNoS_ (SNoLev y).
An exact proof term for the current goal is SNoR_SNoS y Hy z1 Hz1.
We prove the intermediate claim Lz1b: SNo z1.
Apply SNoR_E y Hy z1 Hz1 to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim L1ba: g z0 y = h z0 y.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0.
An exact proof term for the current goal is Hy.
We prove the intermediate claim L1bb: g x z1 = h x z1.
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lz1.
We prove the intermediate claim L1bc: g z0 z1 = h z0 z1.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0.
An exact proof term for the current goal is Lz1b.
rewrite the current goal using L1ba (from left to right).
rewrite the current goal using L1bb (from left to right).
rewrite the current goal using L1bc (from left to right).
Use reflexivity.
We prove the intermediate claim L1c: {g (w 0) y + g x (w 1) + - g (w 0) (w 1)|wSNoL x SNoR y} = {h (w 0) y + h x (w 1) + - h (w 0) (w 1)|wSNoL x SNoR y}.
Apply ReplEq_setprod_ext (SNoL x) (SNoR y) (λw0 w1 ⇒ g w0 y + g x w1 + - g w0 w1) (λw0 w1 ⇒ h w0 y + h x w1 + - h w0 w1) to the current goal.
We will prove w0SNoL x, w1SNoR y, g w0 y + g x w1 + - g w0 w1 = h w0 y + h x w1 + - h w0 w1.
Let w0 be given.
Assume Hw0: w0 SNoL x.
Let w1 be given.
Assume Hw1: w1 SNoR y.
We prove the intermediate claim Lw0: w0 SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoL_SNoS x Hx w0 Hw0.
We prove the intermediate claim Lw1: w1 SNoS_ (SNoLev y).
An exact proof term for the current goal is SNoR_SNoS y Hy w1 Hw1.
We prove the intermediate claim Lw1b: SNo w1.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim L1ca: g w0 y = h w0 y.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0.
An exact proof term for the current goal is Hy.
We prove the intermediate claim L1cb: g x w1 = h x w1.
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lw1.
We prove the intermediate claim L1cc: g w0 w1 = h w0 w1.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0.
An exact proof term for the current goal is Lw1b.
rewrite the current goal using L1ca (from left to right).
rewrite the current goal using L1cb (from left to right).
rewrite the current goal using L1cc (from left to right).
Use reflexivity.
We prove the intermediate claim L1d: {g (z 0) y + g x (z 1) + - g (z 0) (z 1)|zSNoR x SNoL y} = {h (z 0) y + h x (z 1) + - h (z 0) (z 1)|zSNoR x SNoL y}.
Apply ReplEq_setprod_ext (SNoR x) (SNoL y) (λz0 z1 ⇒ g z0 y + g x z1 + - g z0 z1) (λz0 z1 ⇒ h z0 y + h x z1 + - h z0 z1) to the current goal.
We will prove z0SNoR x, z1SNoL y, g z0 y + g x z1 + - g z0 z1 = h z0 y + h x z1 + - h z0 z1.
Let z0 be given.
Assume Hz0: z0 SNoR x.
Let z1 be given.
Assume Hz1: z1 SNoL y.
We prove the intermediate claim Lz0: z0 SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoR_SNoS x Hx z0 Hz0.
We prove the intermediate claim Lz1: z1 SNoS_ (SNoLev y).
An exact proof term for the current goal is SNoL_SNoS y Hy z1 Hz1.
We prove the intermediate claim Lz1b: SNo z1.
Apply SNoL_E y Hy z1 Hz1 to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim L1da: g z0 y = h z0 y.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0.
An exact proof term for the current goal is Hy.
We prove the intermediate claim L1db: g x z1 = h x z1.
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lz1.
We prove the intermediate claim L1dc: g z0 z1 = h z0 z1.
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0.
An exact proof term for the current goal is Lz1b.
rewrite the current goal using L1da (from left to right).
rewrite the current goal using L1db (from left to right).
rewrite the current goal using L1dc (from left to right).
Use reflexivity.
We will prove SNoCut ({g (w 0) y + g x (w 1) + - g (w 0) (w 1)|wSNoL x SNoL y} {g (z 0) y + g x (z 1) + - g (z 0) (z 1)|zSNoR x SNoR y}) ({g (w 0) y + g x (w 1) + - g (w 0) (w 1)|wSNoL x SNoR y} {g (z 0) y + g x (z 1) + - g (z 0) (z 1)|zSNoR x SNoL y}) = SNoCut ({h (w 0) y + h x (w 1) + - h (w 0) (w 1)|wSNoL x SNoL y} {h (z 0) y + h x (z 1) + - h (z 0) (z 1)|zSNoR x SNoR y}) ({h (w 0) y + h x (w 1) + - h (w 0) (w 1)|wSNoL x SNoR y} {h (z 0) y + h x (z 1) + - h (z 0) (z 1)|zSNoR x SNoL y}).
rewrite the current goal using L1a (from left to right).
rewrite the current goal using L1b (from left to right).
rewrite the current goal using L1c (from left to right).
rewrite the current goal using L1d (from left to right).
Use reflexivity.
An exact proof term for the current goal is SNo_rec2_eq F L1.
Theorem. (mul_SNo_eq_2) The following is provable:
∀x y, SNo xSNo y∀p : prop, (∀L R, (∀u, u L(∀q : prop, (w0SNoL x, w1SNoL y, u = w0 * y + x * w1 + - w0 * w1q)(z0SNoR x, z1SNoR y, u = z0 * y + x * z1 + - z0 * z1q)q))(w0SNoL x, w1SNoL y, w0 * y + x * w1 + - w0 * w1 L)(z0SNoR x, z1SNoR y, z0 * y + x * z1 + - z0 * z1 L)(∀u, u R(∀q : prop, (w0SNoL x, z1SNoR y, u = w0 * y + x * z1 + - w0 * z1q)(z0SNoR x, w1SNoL y, u = z0 * y + x * w1 + - z0 * w1q)q))(w0SNoL x, z1SNoR y, w0 * y + x * z1 + - w0 * z1 R)(z0SNoR x, w1SNoL y, z0 * y + x * w1 + - z0 * w1 R)x * y = SNoCut L Rp)p
Proof:
Let x and y be given.
Assume Hx Hy.
Let p be given.
Assume Hp.
We will prove p.
Set Lxy1 to be the term {(w 0) * y + x * (w 1) + - (w 0) * (w 1)|wSNoL x SNoL y}.
Set Lxy2 to be the term {(z 0) * y + x * (z 1) + - (z 0) * (z 1)|zSNoR x SNoR y}.
Set Rxy1 to be the term {(w 0) * y + x * (w 1) + - (w 0) * (w 1)|wSNoL x SNoR y}.
Set Rxy2 to be the term {(z 0) * y + x * (z 1) + - (z 0) * (z 1)|zSNoR x SNoL y}.
Apply Hp (Lxy1 Lxy2) (Rxy1 Rxy2) to the current goal.
Let u be given.
Assume Hu.
Let q be given.
Assume Hq1 Hq2.
Apply binunionE Lxy1 Lxy2 u Hu to the current goal.
Assume Hu1: u Lxy1.
Apply ReplE_impred (SNoL x SNoL y) (λw ⇒ (w 0) * y + x * (w 1) + - (w 0) * (w 1)) u Hu1 to the current goal.
Let w be given.
Assume Hw: w SNoL x SNoL y.
Assume Hw2: u = (w 0) * y + x * (w 1) + - (w 0) * (w 1).
An exact proof term for the current goal is Hq1 (w 0) (ap0_Sigma (SNoL x) (λ_ ⇒ SNoL y) w Hw) (w 1) (ap1_Sigma (SNoL x) (λ_ ⇒ SNoL y) w Hw) Hw2.
Assume Hu2: u Lxy2.
Apply ReplE_impred (SNoR x SNoR y) (λz ⇒ (z 0) * y + x * (z 1) + - (z 0) * (z 1)) u Hu2 to the current goal.
Let z be given.
Assume Hz: z SNoR x SNoR y.
Assume Hz2: u = (z 0) * y + x * (z 1) + - (z 0) * (z 1).
An exact proof term for the current goal is Hq2 (z 0) (ap0_Sigma (SNoR x) (λ_ ⇒ SNoR y) z Hz) (z 1) (ap1_Sigma (SNoR x) (λ_ ⇒ SNoR y) z Hz) Hz2.
Let w0 be given.
Assume Hw0: w0 SNoL x.
Let w1 be given.
Assume Hw1: w1 SNoL y.
We will prove w0 * y + x * w1 + - w0 * w1 Lxy1 Lxy2.
Apply binunionI1 to the current goal.
We will prove w0 * y + x * w1 + - w0 * w1 Lxy1.
Apply tuple_2_0_eq w0 w1 (λu v ⇒ u * y + x * w1 + - u * w1 Lxy1) to the current goal.
We will prove (w0,w1) 0 * y + x * w1 + - (w0,w1) 0 * w1 Lxy1.
Apply tuple_2_1_eq w0 w1 (λu v ⇒ (w0,w1) 0 * y + x * u + - (w0,w1) 0 * u Lxy1) to the current goal.
We will prove (w0,w1) 0 * y + x * (w0,w1) 1 + - (w0,w1) 0 * (w0,w1) 1 Lxy1.
Apply ReplI (SNoL x SNoL y) (λw ⇒ (w 0) * y + x * (w 1) + - (w 0) * (w 1)) (w0,w1) to the current goal.
We will prove (w0,w1) SNoL x SNoL y.
Apply tuple_2_setprod to the current goal.
An exact proof term for the current goal is Hw0.
An exact proof term for the current goal is Hw1.
Let z0 be given.
Assume Hz0: z0 SNoR x.
Let z1 be given.
Assume Hz1: z1 SNoR y.
We will prove z0 * y + x * z1 + - z0 * z1 Lxy1 Lxy2.
Apply binunionI2 to the current goal.
We will prove z0 * y + x * z1 + - z0 * z1 Lxy2.
Apply tuple_2_0_eq z0 z1 (λu v ⇒ u * y + x * z1 + - u * z1 Lxy2) to the current goal.
We will prove (z0,z1) 0 * y + x * z1 + - (z0,z1) 0 * z1 Lxy2.
Apply tuple_2_1_eq z0 z1 (λu v ⇒ (z0,z1) 0 * y + x * u + - (z0,z1) 0 * u Lxy2) to the current goal.
We will prove (z0,z1) 0 * y + x * (z0,z1) 1 + - (z0,z1) 0 * (z0,z1) 1 Lxy2.
Apply ReplI (SNoR x SNoR y) (λz ⇒ (z 0) * y + x * (z 1) + - (z 0) * (z 1)) (z0,z1) to the current goal.
We will prove (z0,z1) SNoR x SNoR y.
Apply tuple_2_setprod to the current goal.
An exact proof term for the current goal is Hz0.
An exact proof term for the current goal is Hz1.
Let u be given.
Assume Hu.
Let q be given.
Assume Hq1 Hq2.
Apply binunionE Rxy1 Rxy2 u Hu to the current goal.
Assume Hu1: u Rxy1.
Apply ReplE_impred (SNoL x SNoR y) (λw ⇒ (w 0) * y + x * (w 1) + - (w 0) * (w 1)) u Hu1 to the current goal.
Let w be given.
Assume Hw: w SNoL x SNoR y.
Assume Hw2: u = (w 0) * y + x * (w 1) + - (w 0) * (w 1).
An exact proof term for the current goal is Hq1 (w 0) (ap0_Sigma (SNoL x) (λ_ ⇒ SNoR y) w Hw) (w 1) (ap1_Sigma (SNoL x) (λ_ ⇒ SNoR y) w Hw) Hw2.
Assume Hu2: u Rxy2.
Apply ReplE_impred (SNoR x SNoL y) (λz ⇒ (z 0) * y + x * (z 1) + - (z 0) * (z 1)) u Hu2 to the current goal.
Let z be given.
Assume Hz: z SNoR x SNoL y.
Assume Hz2: u = (z 0) * y + x * (z 1) + - (z 0) * (z 1).
An exact proof term for the current goal is Hq2 (z 0) (ap0_Sigma (SNoR x) (λ_ ⇒ SNoL y) z Hz) (z 1) (ap1_Sigma (SNoR x) (λ_ ⇒ SNoL y) z Hz) Hz2.
Let w0 be given.
Assume Hw0: w0 SNoL x.
Let z1 be given.
Assume Hz1: z1 SNoR y.
We will prove w0 * y + x * z1 + - w0 * z1 Rxy1 Rxy2.
Apply binunionI1 to the current goal.
We will prove w0 * y + x * z1 + - w0 * z1 Rxy1.
Apply tuple_2_0_eq w0 z1 (λu v ⇒ u * y + x * z1 + - u * z1 Rxy1) to the current goal.
We will prove (w0,z1) 0 * y + x * z1 + - (w0,z1) 0 * z1 Rxy1.
Apply tuple_2_1_eq w0 z1 (λu v ⇒ (w0,z1) 0 * y + x * u + - (w0,z1) 0 * u Rxy1) to the current goal.
We will prove (w0,z1) 0 * y + x * (w0,z1) 1 + - (w0,z1) 0 * (w0,z1) 1 Rxy1.
Apply ReplI (SNoL x SNoR y) (λw ⇒ (w 0) * y + x * (w 1) + - (w 0) * (w 1)) (w0,z1) to the current goal.
We will prove (w0,z1) SNoL x SNoR y.
Apply tuple_2_setprod to the current goal.
An exact proof term for the current goal is Hw0.
An exact proof term for the current goal is Hz1.
Let z0 be given.
Assume Hz0: z0 SNoR x.
Let w1 be given.
Assume Hw1: w1 SNoL y.
We will prove z0 * y + x * w1 + - z0 * w1 Rxy1 Rxy2.
Apply binunionI2 to the current goal.
We will prove z0 * y + x * w1 + - z0 * w1 Rxy2.
Apply tuple_2_0_eq z0 w1 (λu v ⇒ u * y + x * w1 + - u * w1 Rxy2) to the current goal.
We will prove (z0,w1) 0 * y + x * w1 + - (z0,w1) 0 * w1 Rxy2.
Apply tuple_2_1_eq z0 w1 (λu v ⇒ (z0,w1) 0 * y + x * u + - (z0,w1) 0 * u Rxy2) to the current goal.
We will prove (z0,w1) 0 * y + x * (z0,w1) 1 + - (z0,w1) 0 * (z0,w1) 1 Rxy2.
Apply ReplI (SNoR x SNoL y) (λw ⇒ (w 0) * y + x * (w 1) + - (w 0) * (w 1)) (z0,w1) to the current goal.
We will prove (z0,w1) SNoR x SNoL y.
Apply tuple_2_setprod to the current goal.
An exact proof term for the current goal is Hz0.
An exact proof term for the current goal is Hw1.
An exact proof term for the current goal is mul_SNo_eq x Hx y Hy.
Theorem. (mul_SNo_prop_1) The following is provable:
∀x, SNo x∀y, SNo y∀p : prop, (SNo (x * y)(uSNoL x, vSNoL y, u * y + x * v < x * y + u * v)(uSNoR x, vSNoR y, u * y + x * v < x * y + u * v)(uSNoL x, vSNoR y, x * y + u * v < u * y + x * v)(uSNoR x, vSNoL y, x * y + u * v < u * y + x * v)p)p
Proof:
Set P to be the term λx ⇒ ∀y, SNo y∀p : prop, (SNo (x * y)(uSNoL x, vSNoL y, u * y + x * v < x * y + u * v)(uSNoR x, vSNoR y, u * y + x * v < x * y + u * v)(uSNoL x, vSNoR y, x * y + u * v < u * y + x * v)(uSNoR x, vSNoL y, x * y + u * v < u * y + x * v)p)p of type setprop.
We will prove ∀x, SNo xP x.
Apply SNoLev_ind to the current goal.
Let x be given.
Assume Hx: SNo x.
Assume IHx: wSNoS_ (SNoLev x), P w.
Set Q to be the term λy ⇒ ∀p : prop, (SNo (x * y)(uSNoL x, vSNoL y, u * y + x * v < x * y + u * v)(uSNoR x, vSNoR y, u * y + x * v < x * y + u * v)(uSNoL x, vSNoR y, x * y + u * v < u * y + x * v)(uSNoR x, vSNoL y, x * y + u * v < u * y + x * v)p)p of type setprop.
We will prove ∀y, SNo yQ y.
Apply SNoLev_ind to the current goal.
Let y be given.
Assume Hy: SNo y.
Assume IHy: zSNoS_ (SNoLev y), Q z.
Apply mul_SNo_eq_2 x y Hx Hy to the current goal.
Let L and R be given.
Assume HLE HLI1 HLI2 HRE HRI1 HRI2 Hxy.
We prove the intermediate claim LLx: uSNoL x, ∀v, SNo vSNo (u * v).
Let u be given.
Assume Hu.
Let v be given.
Assume Hv.
Apply IHx u (SNoL_SNoS x Hx u Hu) v Hv to the current goal.
Assume H1 _ _ _ _.
An exact proof term for the current goal is H1.
We prove the intermediate claim LRx: uSNoR x, ∀v, SNo vSNo (u * v).
Let u be given.
Assume Hu.
Let v be given.
Assume Hv.
Apply IHx u (SNoR_SNoS x Hx u Hu) v Hv to the current goal.
Assume H1 _ _ _ _.
An exact proof term for the current goal is H1.
We prove the intermediate claim LLxy: uSNoL x, SNo (u * y).
Let u be given.
Assume Hu.
An exact proof term for the current goal is LLx u Hu y Hy.
We prove the intermediate claim LRxy: uSNoR x, SNo (u * y).
Let u be given.
Assume Hu.
An exact proof term for the current goal is LRx u Hu y Hy.
We prove the intermediate claim LxLy: vSNoL y, SNo (x * v).
Let v be given.
Assume Hv.
Apply IHy v (SNoL_SNoS y Hy v Hv) to the current goal.
Assume H1 _ _ _ _.
An exact proof term for the current goal is H1.
We prove the intermediate claim LxRy: vSNoR y, SNo (x * v).
Let v be given.
Assume Hv.
Apply IHy v (SNoR_SNoS y Hy v Hv) to the current goal.
Assume H1 _ _ _ _.
An exact proof term for the current goal is H1.
We prove the intermediate claim LLR1: SNoCutP L R.
We will prove (xL, SNo x) (yR, SNo y) (xL, yR, x < y).
Apply and3I to the current goal.
Let u be given.
Assume Hu.
Apply HLE u Hu to the current goal.
Let w0 be given.
Assume Hw0.
Let w1 be given.
Assume Hw1 Huw.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw1a _ _.
We will prove SNo u.
rewrite the current goal using Huw (from left to right).
We will prove SNo (w0 * y + x * w1 + - w0 * w1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (w0 * y).
An exact proof term for the current goal is LLxy w0 Hw0.
We will prove SNo (x * w1 + - w0 * w1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (x * w1).
An exact proof term for the current goal is LxLy w1 Hw1.
We will prove SNo (- w0 * w1).
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is LLx w0 Hw0 w1 Hw1a.
Let z0 be given.
Assume Hz0.
Let z1 be given.
Assume Hz1 Huz.
Apply SNoR_E y Hy z1 Hz1 to the current goal.
Assume Hz1a _ _.
We will prove SNo u.
rewrite the current goal using Huz (from left to right).
We will prove SNo (z0 * y + x * z1 + - z0 * z1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (z0 * y).
An exact proof term for the current goal is LRxy z0 Hz0.
We will prove SNo (x * z1 + - z0 * z1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (x * z1).
An exact proof term for the current goal is LxRy z1 Hz1.
We will prove SNo (- z0 * z1).
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is LRx z0 Hz0 z1 Hz1a.
Let u be given.
Assume Hu.
Apply HRE u Hu to the current goal.
Let w0 be given.
Assume Hw0.
Let z1 be given.
Assume Hz1 Huw.
Apply SNoR_E y Hy z1 Hz1 to the current goal.
Assume Hz1a _ _.
We will prove SNo u.
rewrite the current goal using Huw (from left to right).
We will prove SNo (w0 * y + x * z1 + - w0 * z1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (w0 * y).
An exact proof term for the current goal is LLxy w0 Hw0.
We will prove SNo (x * z1 + - w0 * z1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (x * z1).
An exact proof term for the current goal is LxRy z1 Hz1.
We will prove SNo (- w0 * z1).
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is LLx w0 Hw0 z1 Hz1a.
Let z0 be given.
Assume Hz0.
Let w1 be given.
Assume Hw1 Huz.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw1a _ _.
We will prove SNo u.
rewrite the current goal using Huz (from left to right).
We will prove SNo (z0 * y + x * w1 + - z0 * w1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (z0 * y).
An exact proof term for the current goal is LRxy z0 Hz0.
We will prove SNo (x * w1 + - z0 * w1).
Apply SNo_add_SNo to the current goal.
We will prove SNo (x * w1).
An exact proof term for the current goal is LxLy w1 Hw1.
We will prove SNo (- z0 * w1).
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is LRx z0 Hz0 w1 Hw1a.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv.
We prove the intermediate claim Luvimp: u0 v0SNoS_ (SNoLev x), u1 v1SNoS_ (SNoLev y), ∀p : prop, (SNo (u0 * y)SNo (x * u1)SNo (u0 * u1)SNo (v0 * y)SNo (x * v1)SNo (v0 * v1)SNo (u0 * v1)SNo (v0 * u1)(u = u0 * y + x * u1 + - u0 * u1v = v0 * y + x * v1 + - v0 * v1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1u < v)p)p.
Let u0 be given.
Assume Hu0.
Let v0 be given.
Assume Hv0.
Let u1 be given.
Assume Hu1.
Let v1 be given.
Assume Hv1.
Apply SNoS_E2 (SNoLev y) (SNoLev_ordinal y Hy) u1 Hu1 to the current goal.
Assume Hu1a: SNoLev u1 SNoLev y.
Assume Hu1b: ordinal (SNoLev u1).
Assume Hu1c: SNo u1.
Assume _.
Apply SNoS_E2 (SNoLev y) (SNoLev_ordinal y Hy) v1 Hv1 to the current goal.
Assume Hv1a: SNoLev v1 SNoLev y.
Assume Hv1b: ordinal (SNoLev v1).
Assume Hv1c: SNo v1.
Assume _.
We prove the intermediate claim Lu0u1: SNo (u0 * u1).
Apply IHx u0 Hu0 u1 Hu1c to the current goal.
Assume IHx0 _ _ _ _.
An exact proof term for the current goal is IHx0.
We prove the intermediate claim Lv0v1: SNo (v0 * v1).
Apply IHx v0 Hv0 v1 Hv1c to the current goal.
Assume IHx0 _ _ _ _.
An exact proof term for the current goal is IHx0.
We prove the intermediate claim Lu0y: SNo (u0 * y).
Apply IHx u0 Hu0 y Hy to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lxu1: SNo (x * u1).
Apply IHy u1 Hu1 to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lv0y: SNo (v0 * y).
Apply IHx v0 Hv0 y Hy to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lxv1: SNo (x * v1).
Apply IHy v1 Hv1 to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lu0v1: SNo (u0 * v1).
Apply IHx u0 Hu0 v1 Hv1c to the current goal.
Assume IHx0 _ _ _ _.
An exact proof term for the current goal is IHx0.
We prove the intermediate claim Lv0u1: SNo (v0 * u1).
Apply IHx v0 Hv0 u1 Hu1c to the current goal.
Assume IHx0 _ _ _ _.
An exact proof term for the current goal is IHx0.
Let p be given.
Assume Hp.
Apply Hp Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 to the current goal.
Assume Hue Hve H1.
We will prove u < v.
rewrite the current goal using Hue (from left to right).
rewrite the current goal using Hve (from left to right).
Apply add_SNo_Lt1_cancel (u0 * y + x * u1 + - u0 * u1) (u0 * u1 + v0 * v1) (v0 * y + x * v1 + - v0 * v1) (SNo_add_SNo (u0 * y) (x * u1 + - u0 * u1) Lu0y (SNo_add_SNo (x * u1) (- u0 * u1) Lxu1 (SNo_minus_SNo (u0 * u1) Lu0u1))) (SNo_add_SNo (u0 * u1) (v0 * v1) Lu0u1 Lv0v1) (SNo_add_SNo (v0 * y) (x * v1 + - v0 * v1) Lv0y (SNo_add_SNo (x * v1) (- v0 * v1) Lxv1 (SNo_minus_SNo (v0 * v1) Lv0v1))) to the current goal.
We prove the intermediate claim L1: (u0 * y + x * u1 + - u0 * u1) + (u0 * u1 + v0 * v1) = u0 * y + x * u1 + v0 * v1.
An exact proof term for the current goal is add_SNo_minus_SNo_prop5 (u0 * y) (x * u1) (u0 * u1) (v0 * v1) Lu0y Lxu1 Lu0u1 Lv0v1.
We prove the intermediate claim L2: (v0 * y + x * v1 + - v0 * v1) + (u0 * u1 + v0 * v1) = v0 * y + x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (u0 * u1) (v0 * v1) Lu0u1 Lv0v1 (from left to right).
An exact proof term for the current goal is add_SNo_minus_SNo_prop5 (v0 * y) (x * v1) (v0 * v1) (u0 * u1) Lv0y Lxv1 Lv0v1 Lu0u1.
rewrite the current goal using L1 (from left to right).
rewrite the current goal using L2 (from left to right).
An exact proof term for the current goal is H1.
Apply HLE u Hu to the current goal.
Let u0 be given.
Assume Hu0.
Let u1 be given.
Assume Hu1 Hue.
Apply SNoL_E x Hx u0 Hu0 to the current goal.
Assume Hu0a: SNo u0.
Assume Hu0b: SNoLev u0 SNoLev x.
Assume Hu0c: u0 < x.
Apply SNoL_E y Hy u1 Hu1 to the current goal.
Assume Hu1a: SNo u1.
Assume Hu1b: SNoLev u1 SNoLev y.
Assume Hu1c: u1 < y.
Apply HRE v Hv to the current goal.
Let v0 be given.
Assume Hv0.
Let v1 be given.
Assume Hv1 Hve.
Apply SNoL_E x Hx v0 Hv0 to the current goal.
Assume Hv0a: SNo v0.
Assume Hv0b: SNoLev v0 SNoLev x.
Assume Hv0c: v0 < x.
Apply SNoR_E y Hy v1 Hv1 to the current goal.
Assume Hv1a: SNo v1.
Assume Hv1b: SNoLev v1 SNoLev y.
Assume Hv1c: y < v1.
Apply Luvimp u0 (SNoL_SNoS x Hx u0 Hu0) v0 (SNoL_SNoS x Hx v0 Hv0) u1 (SNoL_SNoS y Hy u1 Hu1) v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv.
Apply Luv Hue Hve to the current goal.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
We prove the intermediate claim Lu1v1lt: u1 < v1.
An exact proof term for the current goal is SNoLt_tra u1 y v1 Hu1a Hy Hv1a Hu1c Hv1c.
We prove the intermediate claim L3: zSNoL x, x * u1 + z * v1 < z * u1 + x * v1.
Let z be given.
Assume Hz.
We prove the intermediate claim Lzu1: SNo (z * u1).
An exact proof term for the current goal is LLx z Hz u1 Hu1a.
We prove the intermediate claim Lzv1: SNo (z * v1).
An exact proof term for the current goal is LLx z Hz v1 Hv1a.
Apply SNoLt_SNoL_or_SNoR_impred u1 v1 Hu1a Hv1a Lu1v1lt to the current goal.
Let w be given.
Assume Hwv1: w SNoL v1.
Assume Hwu1: w SNoR u1.
Apply SNoR_E u1 Hu1a w Hwu1 to the current goal.
Assume Hwu1a: SNo w.
Assume Hwu1b: SNoLev w SNoLev u1.
Assume Hwu1c: u1 < w.
We prove the intermediate claim LwSy: w SNoS_ (SNoLev y).
Apply SNoS_I2 w y Hwu1a Hy to the current goal.
We will prove SNoLev w SNoLev y.
An exact proof term for the current goal is ordinal_TransSet (SNoLev y) (SNoLev_ordinal y Hy) (SNoLev u1) Hu1b (SNoLev w) Hwu1b.
We prove the intermediate claim Lzw: SNo (z * w).
An exact proof term for the current goal is LLx z Hz w Hwu1a.
We prove the intermediate claim Lxw: SNo (x * w).
Apply IHy w LwSy to the current goal.
Assume H1 _ _ _ _.
An exact proof term for the current goal is H1.
We prove the intermediate claim L3a: z * v1 + x * w < x * v1 + z * w.
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume _ IHy1 _ _ _.
An exact proof term for the current goal is IHy1 z Hz w Hwv1.
We prove the intermediate claim L3b: x * u1 + z * w < z * u1 + x * w.
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ IHy3 _.
An exact proof term for the current goal is IHy3 z Hz w Hwu1.
An exact proof term for the current goal is add_SNo_Lt_subprop2 (x * u1) (z * v1) (z * u1) (x * v1) (z * w) (x * w) Lxu1 Lzv1 Lzu1 Lxv1 Lzw Lxw L3b L3a.
Assume Hu1v1: u1 SNoL v1.
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume _ IHy1 _ _ _.
rewrite the current goal using add_SNo_com (x * u1) (z * v1) Lxu1 Lzv1 (from left to right).
rewrite the current goal using add_SNo_com (z * u1) (x * v1) Lzu1 Lxv1 (from left to right).
We will prove z * v1 + x * u1 < x * v1 + z * u1.
An exact proof term for the current goal is IHy1 z Hz u1 Hu1v1.
Assume Hv1u1: v1 SNoR u1.
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ IHy3 _.
An exact proof term for the current goal is IHy3 z Hz v1 Hv1u1.
We prove the intermediate claim L3u0: x * u1 + u0 * v1 < u0 * u1 + x * v1.
An exact proof term for the current goal is L3 u0 Hu0.
We prove the intermediate claim L3v0: x * u1 + v0 * v1 < v0 * u1 + x * v1.
An exact proof term for the current goal is L3 v0 Hv0.
We prove the intermediate claim L3u0imp: u0 * y + v0 * v1 < v0 * y + u0 * v1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
Assume H1.
Apply add_SNo_Lt_subprop3a (u0 * y) (x * u1) (v0 * v1) (v0 * y) (x * v1 + u0 * u1) (u0 * v1) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) Lu0v1 H1 to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0.
We prove the intermediate claim L3v0imp: u0 * y + v0 * u1 < v0 * y + u0 * u1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
Assume H1.
Apply add_SNo_Lt_subprop3b (u0 * y) (x * u1 + v0 * v1) (v0 * y) (x * v1) (u0 * u1) (v0 * u1) Lu0y (SNo_add_SNo (x * u1) (v0 * v1) Lxu1 Lv0v1) Lv0y Lxv1 Lu0u1 Lv0u1 to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
An exact proof term for the current goal is H1.
We will prove x * u1 + v0 * v1 < v0 * u1 + x * v1.
An exact proof term for the current goal is L3v0.
Apply SNoL_or_SNoR_impred v0 u0 Hv0a Hu0a to the current goal.
Assume Hv0u0: v0 = u0.
rewrite the current goal using Hv0u0 (from left to right).
We will prove u0 * y + (x * u1 + u0 * v1) < u0 * y + (x * v1 + u0 * u1).
Apply add_SNo_Lt2 (u0 * y) (x * u1 + u0 * v1) (x * v1 + u0 * u1) Lu0y (SNo_add_SNo (x * u1) (u0 * v1) Lxu1 Lu0v1) (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0.
Let z be given.
Assume Hzu0: z SNoL u0.
Assume Hzv0: z SNoR v0.
We prove the intermediate claim L4: u0 * y + z * v1 < z * y + u0 * v1.
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ IHx3 _.
An exact proof term for the current goal is IHx3 z Hzu0 v1 Hv1.
We prove the intermediate claim L5: z * y + v0 * v1 < v0 * y + z * v1.
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ _ IHx2 _ _.
An exact proof term for the current goal is IHx2 z Hzv0 v1 Hv1.
We prove the intermediate claim Lz: z SNoL x.
Apply SNoL_E u0 Hu0a z Hzu0 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev u0.
Assume Hzc: z < u0.
Apply SNoL_I x Hx z Hza to the current goal.
We will prove SNoLev z SNoLev x.
An exact proof term for the current goal is ordinal_TransSet (SNoLev x) (SNoLev_ordinal x Hx) (SNoLev u0) Hu0b (SNoLev z) Hzb.
We will prove z < x.
An exact proof term for the current goal is SNoLt_tra z u0 x Hza Hu0a Hx Hzc Hu0c.
Apply add_SNo_Lt_subprop3c (u0 * y) (x * u1) (v0 * v1) (v0 * y) (x * v1 + u0 * u1) (z * v1) (z * y) (u0 * v1) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) (LLx z Lz v1 Hv1a) (LLxy z Lz) Lu0v1 to the current goal.
An exact proof term for the current goal is L4.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0.
An exact proof term for the current goal is L5.
Assume Hv0u0: v0 SNoL u0.
Apply L3u0imp to the current goal.
We will prove u0 * y + v0 * v1 < v0 * y + u0 * v1.
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ IHx3 _.
An exact proof term for the current goal is IHx3 v0 Hv0u0 v1 Hv1.
Assume Hu0v0: u0 SNoR v0.
Apply L3u0imp to the current goal.
We will prove u0 * y + v0 * v1 < v0 * y + u0 * v1.
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ _ IHx2 _ _.
An exact proof term for the current goal is IHx2 u0 Hu0v0 v1 Hv1.
Let z be given.
Assume Hzu0: z SNoR u0.
Assume Hzv0: z SNoL v0.
We prove the intermediate claim L6: z * y + v0 * u1 < v0 * y + z * u1.
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ IHx1 _ _ _.
An exact proof term for the current goal is IHx1 z Hzv0 u1 Hu1.
We prove the intermediate claim L7: u0 * y + z * u1 < z * y + u0 * u1.
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ _ IHx4.
An exact proof term for the current goal is IHx4 z Hzu0 u1 Hu1.
We prove the intermediate claim Lz: z SNoL x.
Apply SNoL_E v0 Hv0a z Hzv0 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev v0.
Assume Hzc: z < v0.
Apply SNoL_I x Hx z Hza to the current goal.
We will prove SNoLev z SNoLev x.
An exact proof term for the current goal is ordinal_TransSet (SNoLev x) (SNoLev_ordinal x Hx) (SNoLev v0) Hv0b (SNoLev z) Hzb.
We will prove z < x.
An exact proof term for the current goal is SNoLt_tra z v0 x Hza Hv0a Hx Hzc Hv0c.
An exact proof term for the current goal is add_SNo_Lt_subprop3d (u0 * y) (x * u1 + v0 * v1) (v0 * y) (x * v1) (u0 * u1) (z * u1) (z * y) (v0 * u1) Lu0y (SNo_add_SNo (x * u1) (v0 * v1) Lxu1 Lv0v1) Lv0y Lxv1 Lu0u1 (LLx z Lz u1 Hu1a) (LLxy z Lz) Lv0u1 L7 L3v0 L6.
Assume Hv0u0: v0 SNoR u0.
Apply L3v0imp to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ _ IHx4.
An exact proof term for the current goal is IHx4 v0 Hv0u0 u1 Hu1.
Assume Hu0v0: u0 SNoL v0.
Apply L3v0imp to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ IHx1 _ _ _.
An exact proof term for the current goal is IHx1 u0 Hu0v0 u1 Hu1.
Let v0 be given.
Assume Hv0.
Let v1 be given.
Assume Hv1 Hve.
Apply SNoR_E x Hx v0 Hv0 to the current goal.
Assume Hv0a: SNo v0.
Assume Hv0b: SNoLev v0 SNoLev x.
Assume Hv0c: x < v0.
Apply SNoL_E y Hy v1 Hv1 to the current goal.
Assume Hv1a: SNo v1.
Assume Hv1b: SNoLev v1 SNoLev y.
Assume Hv1c: v1 < y.
Apply Luvimp u0 (SNoL_SNoS x Hx u0 Hu0) v0 (SNoR_SNoS x Hx v0 Hv0) u1 (SNoL_SNoS y Hy u1 Hu1) v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv.
Apply Luv Hue Hve to the current goal.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
We prove the intermediate claim Lu0v0lt: u0 < v0.
An exact proof term for the current goal is SNoLt_tra u0 x v0 Hu0a Hx Hv0a Hu0c Hv0c.
We prove the intermediate claim L3: zSNoL y, u0 * y + v0 * z < v0 * y + u0 * z.
Let z be given.
Assume Hz.
Apply SNoL_E y Hy z Hz to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev y.
Assume Hzc: z < y.
We prove the intermediate claim Lu0z: SNo (u0 * z).
An exact proof term for the current goal is LLx u0 Hu0 z Hza.
We prove the intermediate claim Lv0z: SNo (v0 * z).
An exact proof term for the current goal is LRx v0 Hv0 z Hza.
Apply SNoLt_SNoL_or_SNoR_impred u0 v0 Hu0a Hv0a Lu0v0lt to the current goal.
Let w be given.
Assume Hwv0: w SNoL v0.
Assume Hwu0: w SNoR u0.
Apply SNoR_E u0 Hu0a w Hwu0 to the current goal.
Assume Hwu0a: SNo w.
Assume Hwu0b: SNoLev w SNoLev u0.
Assume Hwu0c: u0 < w.
We prove the intermediate claim LLevwLevx: SNoLev w SNoLev x.
An exact proof term for the current goal is ordinal_TransSet (SNoLev x) (SNoLev_ordinal x Hx) (SNoLev u0) Hu0b (SNoLev w) Hwu0b.
We prove the intermediate claim LwSx: w SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 w x Hwu0a Hx LLevwLevx.
We prove the intermediate claim Lwz: SNo (w * z).
Apply IHx w LwSx z Hza to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lwy: SNo (w * y).
Apply IHx w LwSx y Hy to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim L3a: u0 * y + w * z < u0 * z + w * y.
rewrite the current goal using add_SNo_com (u0 * z) (w * y) Lu0z Lwy (from left to right).
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ _ IHx4.
An exact proof term for the current goal is IHx4 w Hwu0 z Hz.
We prove the intermediate claim L3b: v0 * z + w * y < v0 * y + w * z.
rewrite the current goal using add_SNo_com (v0 * z) (w * y) Lv0z Lwy (from left to right).
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ IHx1 _ _ _.
An exact proof term for the current goal is IHx1 w Hwv0 z Hz.
rewrite the current goal using add_SNo_com (u0 * z) (v0 * y) Lu0z Lv0y (from right to left).
An exact proof term for the current goal is add_SNo_Lt_subprop2 (u0 * y) (v0 * z) (u0 * z) (v0 * y) (w * z) (w * y) Lu0y Lv0z Lu0z Lv0y Lwz Lwy L3a L3b.
Assume Hu0v0: u0 SNoL v0.
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ IHx1 _ _ _.
An exact proof term for the current goal is IHx1 u0 Hu0v0 z Hz.
Assume Hv0u0: v0 SNoR u0.
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ _ IHx4.
An exact proof term for the current goal is IHx4 v0 Hv0u0 z Hz.
We prove the intermediate claim L3u1: u0 * y + v0 * u1 < v0 * y + u0 * u1.
An exact proof term for the current goal is L3 u1 Hu1.
We prove the intermediate claim L3v1: u0 * y + v0 * v1 < v0 * y + u0 * v1.
An exact proof term for the current goal is L3 v1 Hv1.
We prove the intermediate claim L3u1imp: x * u1 + v0 * v1 < v0 * u1 + x * v1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
An exact proof term for the current goal is add_SNo_Lt_subprop3b (u0 * y) (x * u1 + v0 * v1) (v0 * y) (x * v1) (u0 * u1) (v0 * u1) Lu0y (SNo_add_SNo (x * u1) (v0 * v1) Lxu1 Lv0v1) Lv0y Lxv1 Lu0u1 Lv0u1 L3u1.
We prove the intermediate claim L3v1imp: x * u1 + u0 * v1 < x * v1 + u0 * u1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
An exact proof term for the current goal is add_SNo_Lt_subprop3a (u0 * y) (x * u1) (v0 * v1) (v0 * y) (x * v1 + u0 * u1) (u0 * v1) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) Lu0v1 L3v1.
Apply SNoL_or_SNoR_impred v1 u1 Hv1a Hu1a to the current goal.
Assume Hv1u1: v1 = u1.
rewrite the current goal using Hv1u1 (from left to right).
We will prove u0 * y + x * u1 + v0 * u1 < v0 * y + x * u1 + u0 * u1.
rewrite the current goal using add_SNo_com_3_0_1 (u0 * y) (x * u1) (v0 * u1) Lu0y Lxu1 Lv0u1 (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (v0 * y) (x * u1) (u0 * u1) Lv0y Lxu1 Lu0u1 (from left to right).
Apply add_SNo_Lt2 (x * u1) (u0 * y + v0 * u1) (v0 * y + u0 * u1) Lxu1 (SNo_add_SNo (u0 * y) (v0 * u1) Lu0y Lv0u1) (SNo_add_SNo (v0 * y) (u0 * u1) Lv0y Lu0u1) to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
An exact proof term for the current goal is L3u1.
Let z be given.
Assume Hzu1: z SNoL u1.
Assume Hzv1: z SNoR v1.
Apply SNoL_E u1 Hu1a z Hzu1 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev u1.
Assume Hzc: z < u1.
We prove the intermediate claim Lz: z SNoL y.
Apply SNoL_I y Hy z Hza to the current goal.
We will prove SNoLev z SNoLev y.
An exact proof term for the current goal is ordinal_TransSet (SNoLev y) (SNoLev_ordinal y Hy) (SNoLev u1) Hu1b (SNoLev z) Hzb.
We will prove z < y.
An exact proof term for the current goal is SNoLt_tra z u1 y Hza Hu1a Hy Hzc Hu1c.
We prove the intermediate claim L4: x * u1 + v0 * z < x * z + v0 * u1.
rewrite the current goal using add_SNo_com (x * z) (v0 * u1) (LxLy z Lz) Lv0u1 (from left to right).
We will prove x * u1 + v0 * z < v0 * u1 + x * z.
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ _ IHy4.
An exact proof term for the current goal is IHy4 v0 Hv0 z Hzu1.
We prove the intermediate claim L5: x * z + v0 * v1 < x * v1 + v0 * z.
rewrite the current goal using add_SNo_com (x * z) (v0 * v1) (LxLy z Lz) Lv0v1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume _ _ IHy2 _ _.
An exact proof term for the current goal is IHy2 v0 Hv0 z Hzv1.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com_3_0_1 (u0 * y) (x * u1) (v0 * v1) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (v0 * y) (x * v1) (u0 * u1) Lv0y Lxv1 Lu0u1 (from left to right).
We will prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1.
An exact proof term for the current goal is add_SNo_Lt_subprop3c (x * u1) (u0 * y) (v0 * v1) (x * v1) (v0 * y + u0 * u1) (v0 * z) (x * z) (v0 * u1) Lxu1 Lu0y Lv0v1 Lxv1 (SNo_add_SNo (v0 * y) (u0 * u1) Lv0y Lu0u1) (LRx v0 Hv0 z Hza) (LxLy z Lz) Lv0u1 L4 L3u1 L5.
Assume Hv1u1: v1 SNoL u1.
Apply L3u1imp to the current goal.
We will prove x * u1 + v0 * v1 < v0 * u1 + x * v1.
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ _ IHy4.
An exact proof term for the current goal is IHy4 v0 Hv0 v1 Hv1u1.
Assume Hu1v1: u1 SNoR v1.
Apply L3u1imp to the current goal.
We will prove x * u1 + v0 * v1 < v0 * u1 + x * v1.
rewrite the current goal using add_SNo_com (x * u1) (v0 * v1) Lxu1 Lv0v1 (from left to right).
rewrite the current goal using add_SNo_com (v0 * u1) (x * v1) Lv0u1 Lxv1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume _ _ IHy2 _ _.
An exact proof term for the current goal is IHy2 v0 Hv0 u1 Hu1v1.
Let z be given.
Assume Hzu1: z SNoR u1.
Assume Hzv1: z SNoL v1.
Apply SNoL_E v1 Hv1a z Hzv1 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev v1.
Assume Hzc: z < v1.
We prove the intermediate claim Lz: z SNoL y.
Apply SNoL_I y Hy z Hza to the current goal.
We will prove SNoLev z SNoLev y.
An exact proof term for the current goal is ordinal_TransSet (SNoLev y) (SNoLev_ordinal y Hy) (SNoLev v1) Hv1b (SNoLev z) Hzb.
We will prove z < y.
An exact proof term for the current goal is SNoLt_tra z v1 y Hza Hv1a Hy Hzc Hv1c.
We prove the intermediate claim L6: x * u1 + u0 * z < x * z + u0 * u1.
rewrite the current goal using add_SNo_com (x * z) (u0 * u1) (LxLy z Lz) Lu0u1 (from left to right).
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ IHy3 _.
An exact proof term for the current goal is IHy3 u0 Hu0 z Hzu1.
We prove the intermediate claim L7: x * z + u0 * v1 < x * v1 + u0 * z.
rewrite the current goal using add_SNo_com (x * z) (u0 * v1) (LxLy z Lz) Lu0v1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume _ IHy1 _ _ _.
An exact proof term for the current goal is IHy1 u0 Hu0 z Hzv1.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com_3_0_1 (u0 * y) (x * u1) (v0 * v1) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (v0 * y) (x * v1) (u0 * u1) Lv0y Lxv1 Lu0u1 (from left to right).
We will prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1.
Apply add_SNo_Lt_subprop3d (x * u1) (u0 * y + v0 * v1) (x * v1) (v0 * y) (u0 * u1) (u0 * z) (x * z) (u0 * v1) Lxu1 (SNo_add_SNo (u0 * y) (v0 * v1) Lu0y Lv0v1) Lxv1 Lv0y Lu0u1 (LLx u0 Hu0 z Hza) (LxLy z Lz) Lu0v1 L6 to the current goal.
We will prove u0 * y + v0 * v1 < u0 * v1 + v0 * y.
rewrite the current goal using add_SNo_com (u0 * v1) (v0 * y) Lu0v1 Lv0y (from left to right).
An exact proof term for the current goal is L3v1.
An exact proof term for the current goal is L7.
Assume Hv1u1: v1 SNoR u1.
Apply L3v1imp to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ IHy3 _.
An exact proof term for the current goal is IHy3 u0 Hu0 v1 Hv1u1.
Assume Hu1v1: u1 SNoL v1.
Apply L3v1imp to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * u1) (u0 * v1) Lxu1 Lu0v1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume _ IHy1 _ _ _.
An exact proof term for the current goal is IHy1 u0 Hu0 u1 Hu1v1.
Let u0 be given.
Assume Hu0.
Let u1 be given.
Assume Hu1 Hue.
Apply SNoR_E x Hx u0 Hu0 to the current goal.
Assume Hu0a: SNo u0.
Assume Hu0b: SNoLev u0 SNoLev x.
Assume Hu0c: x < u0.
Apply SNoR_E y Hy u1 Hu1 to the current goal.
Assume Hu1a: SNo u1.
Assume Hu1b: SNoLev u1 SNoLev y.
Assume Hu1c: y < u1.
Apply HRE v Hv to the current goal.
Let v0 be given.
Assume Hv0.
Let v1 be given.
Assume Hv1 Hve.
Apply SNoL_E x Hx v0 Hv0 to the current goal.
Assume Hv0a: SNo v0.
Assume Hv0b: SNoLev v0 SNoLev x.
Assume Hv0c: v0 < x.
Apply SNoR_E y Hy v1 Hv1 to the current goal.
Assume Hv1a: SNo v1.
Assume Hv1b: SNoLev v1 SNoLev y.
Assume Hv1c: y < v1.
Apply Luvimp u0 (SNoR_SNoS x Hx u0 Hu0) v0 (SNoL_SNoS x Hx v0 Hv0) u1 (SNoR_SNoS y Hy u1 Hu1) v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv.
Apply Luv Hue Hve to the current goal.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
We prove the intermediate claim Lv0u0lt: v0 < u0.
An exact proof term for the current goal is SNoLt_tra v0 x u0 Hv0a Hx Hu0a Hv0c Hu0c.
We prove the intermediate claim L3: zSNoR y, u0 * y + v0 * z < v0 * y + u0 * z.
Let z be given.
Assume Hz.
Apply SNoR_E y Hy z Hz to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev y.
Assume Hzc: y < z.
We prove the intermediate claim Lu0z: SNo (u0 * z).
An exact proof term for the current goal is LRx u0 Hu0 z Hza.
We prove the intermediate claim Lv0z: SNo (v0 * z).
An exact proof term for the current goal is LLx v0 Hv0 z Hza.
Apply SNoLt_SNoL_or_SNoR_impred v0 u0 Hv0a Hu0a Lv0u0lt to the current goal.
Let w be given.
Assume Hwu0: w SNoL u0.
Assume Hwv0: w SNoR v0.
Apply SNoL_E u0 Hu0a w Hwu0 to the current goal.
Assume Hwu0a: SNo w.
Assume Hwu0b: SNoLev w SNoLev u0.
Assume Hwu0c: w < u0.
We prove the intermediate claim LLevwLevx: SNoLev w SNoLev x.
An exact proof term for the current goal is ordinal_TransSet (SNoLev x) (SNoLev_ordinal x Hx) (SNoLev u0) Hu0b (SNoLev w) Hwu0b.
We prove the intermediate claim LwSx: w SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 w x Hwu0a Hx LLevwLevx.
We prove the intermediate claim Lwz: SNo (w * z).
Apply IHx w LwSx z Hza to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lwy: SNo (w * y).
Apply IHx w LwSx y Hy to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim L3a: u0 * y + w * z < u0 * z + w * y.
rewrite the current goal using add_SNo_com (u0 * z) (w * y) Lu0z Lwy (from left to right).
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ IHx3 _.
An exact proof term for the current goal is IHx3 w Hwu0 z Hz.
We prove the intermediate claim L3b: v0 * z + w * y < v0 * y + w * z.
rewrite the current goal using add_SNo_com (v0 * z) (w * y) Lv0z Lwy (from left to right).
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ _ IHx2 _ _.
An exact proof term for the current goal is IHx2 w Hwv0 z Hz.
rewrite the current goal using add_SNo_com (v0 * y) (u0 * z) Lv0y Lu0z (from left to right).
An exact proof term for the current goal is add_SNo_Lt_subprop2 (u0 * y) (v0 * z) (u0 * z) (v0 * y) (w * z) (w * y) Lu0y Lv0z Lu0z Lv0y Lwz Lwy L3a L3b.
Assume Hv0u0: v0 SNoL u0.
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ IHx3 _.
An exact proof term for the current goal is IHx3 v0 Hv0u0 z Hz.
Assume Hu0v0: u0 SNoR v0.
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ _ IHx2 _ _.
An exact proof term for the current goal is IHx2 u0 Hu0v0 z Hz.
We prove the intermediate claim L3u1: u0 * y + v0 * u1 < v0 * y + u0 * u1.
An exact proof term for the current goal is L3 u1 Hu1.
We prove the intermediate claim L3v1: u0 * y + v0 * v1 < v0 * y + u0 * v1.
An exact proof term for the current goal is L3 v1 Hv1.
We prove the intermediate claim L3u1imp: x * u1 + v0 * v1 < v0 * u1 + x * v1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
An exact proof term for the current goal is add_SNo_Lt_subprop3b (u0 * y) (x * u1 + v0 * v1) (v0 * y) (x * v1) (u0 * u1) (v0 * u1) Lu0y (SNo_add_SNo (x * u1) (v0 * v1) Lxu1 Lv0v1) Lv0y Lxv1 Lu0u1 Lv0u1 L3u1.
We prove the intermediate claim L3v1imp: x * u1 + u0 * v1 < x * v1 + u0 * u1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
An exact proof term for the current goal is add_SNo_Lt_subprop3a (u0 * y) (x * u1) (v0 * v1) (v0 * y) (x * v1 + u0 * u1) (u0 * v1) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) Lu0v1 L3v1.
Apply SNoL_or_SNoR_impred v1 u1 Hv1a Hu1a to the current goal.
Assume Hv1u1: v1 = u1.
rewrite the current goal using Hv1u1 (from left to right).
We will prove u0 * y + x * u1 + v0 * u1 < v0 * y + x * u1 + u0 * u1.
rewrite the current goal using add_SNo_com_3_0_1 (u0 * y) (x * u1) (v0 * u1) Lu0y Lxu1 Lv0u1 (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (v0 * y) (x * u1) (u0 * u1) Lv0y Lxu1 Lu0u1 (from left to right).
Apply add_SNo_Lt2 (x * u1) (u0 * y + v0 * u1) (v0 * y + u0 * u1) Lxu1 (SNo_add_SNo (u0 * y) (v0 * u1) Lu0y Lv0u1) (SNo_add_SNo (v0 * y) (u0 * u1) Lv0y Lu0u1) to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
An exact proof term for the current goal is L3u1.
Let z be given.
Assume Hzu1: z SNoL u1.
Assume Hzv1: z SNoR v1.
Apply SNoR_E v1 Hv1a z Hzv1 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev v1.
Assume Hzc: v1 < z.
We prove the intermediate claim Lz: z SNoR y.
Apply SNoR_I y Hy z Hza to the current goal.
We will prove SNoLev z SNoLev y.
An exact proof term for the current goal is ordinal_TransSet (SNoLev y) (SNoLev_ordinal y Hy) (SNoLev v1) Hv1b (SNoLev z) Hzb.
We will prove y < z.
An exact proof term for the current goal is SNoLt_tra y v1 z Hy Hv1a Hza Hv1c Hzc.
We prove the intermediate claim L4: x * u1 + u0 * z < x * z + u0 * u1.
rewrite the current goal using add_SNo_com (x * z) (u0 * u1) (LxRy z Lz) Lu0u1 (from left to right).
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ _ IHy4.
An exact proof term for the current goal is IHy4 u0 Hu0 z Hzu1.
We prove the intermediate claim L5: x * z + u0 * v1 < x * v1 + u0 * z.
rewrite the current goal using add_SNo_com (x * z) (u0 * v1) (LxRy z Lz) Lu0v1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume _ _ IHy2 _ _.
An exact proof term for the current goal is IHy2 u0 Hu0 z Hzv1.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com_3_0_1 (u0 * y) (x * u1) (v0 * v1) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (v0 * y) (x * v1) (u0 * u1) Lv0y Lxv1 Lu0u1 (from left to right).
We will prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1.
Apply add_SNo_Lt_subprop3d (x * u1) (u0 * y + v0 * v1) (x * v1) (v0 * y) (u0 * u1) (u0 * z) (x * z) (u0 * v1) Lxu1 (SNo_add_SNo (u0 * y) (v0 * v1) Lu0y Lv0v1) Lxv1 Lv0y Lu0u1 (LRx u0 Hu0 z Hza) (LxRy z Lz) Lu0v1 L4 to the current goal.
We will prove u0 * y + v0 * v1 < u0 * v1 + v0 * y.
rewrite the current goal using add_SNo_com (u0 * v1) (v0 * y) Lu0v1 Lv0y (from left to right).
An exact proof term for the current goal is L3v1.
An exact proof term for the current goal is L5.
Assume Hv1u1: v1 SNoL u1.
Apply L3v1imp to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
We will prove x * u1 + u0 * v1 < u0 * u1 + x * v1.
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ _ IHy4.
An exact proof term for the current goal is IHy4 u0 Hu0 v1 Hv1u1.
Assume Hu1v1: u1 SNoR v1.
Apply L3v1imp to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * u1) (u0 * v1) Lxu1 Lu0v1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume _ _ IHy2 _ _.
An exact proof term for the current goal is IHy2 u0 Hu0 u1 Hu1v1.
Let z be given.
Assume Hzu1: z SNoR u1.
Assume Hzv1: z SNoL v1.
Apply SNoR_E u1 Hu1a z Hzu1 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev u1.
Assume Hzc: u1 < z.
We prove the intermediate claim Lz: z SNoR y.
Apply SNoR_I y Hy z Hza to the current goal.
We will prove SNoLev z SNoLev y.
An exact proof term for the current goal is ordinal_TransSet (SNoLev y) (SNoLev_ordinal y Hy) (SNoLev u1) Hu1b (SNoLev z) Hzb.
We will prove y < z.
An exact proof term for the current goal is SNoLt_tra y u1 z Hy Hu1a Hza Hu1c Hzc.
We prove the intermediate claim L6: x * u1 + v0 * z < x * z + v0 * u1.
rewrite the current goal using add_SNo_com (x * z) (v0 * u1) (LxRy z Lz) Lv0u1 (from left to right).
We will prove x * u1 + v0 * z < v0 * u1 + x * z.
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ IHy3 _.
An exact proof term for the current goal is IHy3 v0 Hv0 z Hzu1.
We prove the intermediate claim L7: x * z + v0 * v1 < x * v1 + v0 * z.
rewrite the current goal using add_SNo_com (x * z) (v0 * v1) (LxRy z Lz) Lv0v1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume _ IHy1 _ _ _.
An exact proof term for the current goal is IHy1 v0 Hv0 z Hzv1.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com_3_0_1 (u0 * y) (x * u1) (v0 * v1) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (v0 * y) (x * v1) (u0 * u1) Lv0y Lxv1 Lu0u1 (from left to right).
We will prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1.
An exact proof term for the current goal is add_SNo_Lt_subprop3c (x * u1) (u0 * y) (v0 * v1) (x * v1) (v0 * y + u0 * u1) (v0 * z) (x * z) (v0 * u1) Lxu1 Lu0y Lv0v1 Lxv1 (SNo_add_SNo (v0 * y) (u0 * u1) Lv0y Lu0u1) (LLx v0 Hv0 z Hza) (LxRy z Lz) Lv0u1 L6 L3u1 L7.
Assume Hv1u1: v1 SNoR u1.
Apply L3u1imp to the current goal.
We will prove x * u1 + v0 * v1 < v0 * u1 + x * v1.
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ IHy3 _.
An exact proof term for the current goal is IHy3 v0 Hv0 v1 Hv1u1.
Assume Hu1v1: u1 SNoL v1.
Apply L3u1imp to the current goal.
We will prove x * u1 + v0 * v1 < v0 * u1 + x * v1.
rewrite the current goal using add_SNo_com (x * u1) (v0 * v1) Lxu1 Lv0v1 (from left to right).
rewrite the current goal using add_SNo_com (v0 * u1) (x * v1) Lv0u1 Lxv1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1) to the current goal.
Assume _ IHy1 _ _ _.
An exact proof term for the current goal is IHy1 v0 Hv0 u1 Hu1v1.
Let v0 be given.
Assume Hv0.
Let v1 be given.
Assume Hv1 Hve.
Apply SNoR_E x Hx v0 Hv0 to the current goal.
Assume Hv0a: SNo v0.
Assume Hv0b: SNoLev v0 SNoLev x.
Assume Hv0c: x < v0.
Apply SNoL_E y Hy v1 Hv1 to the current goal.
Assume Hv1a: SNo v1.
Assume Hv1b: SNoLev v1 SNoLev y.
Assume Hv1c: v1 < y.
Apply Luvimp u0 (SNoR_SNoS x Hx u0 Hu0) v0 (SNoR_SNoS x Hx v0 Hv0) u1 (SNoR_SNoS y Hy u1 Hu1) v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv.
Apply Luv Hue Hve to the current goal.
We will prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
We prove the intermediate claim Lv1u1lt: v1 < u1.
An exact proof term for the current goal is SNoLt_tra v1 y u1 Hv1a Hy Hu1a Hv1c Hu1c.
We prove the intermediate claim L3: zSNoR x, x * u1 + z * v1 < z * u1 + x * v1.
Let z be given.
Assume Hz.
We prove the intermediate claim Lzu1: SNo (z * u1).
An exact proof term for the current goal is LRx z Hz u1 Hu1a.
We prove the intermediate claim Lzv1: SNo (z * v1).
An exact proof term for the current goal is LRx z Hz v1 Hv1a.
Apply SNoLt_SNoL_or_SNoR_impred v1 u1 Hv1a Hu1a Lv1u1lt to the current goal.
Let w be given.
Assume Hwu1: w SNoL u1.
Assume Hwv1: w SNoR v1.
Apply SNoL_E u1 Hu1a w Hwu1 to the current goal.
Assume Hwu1a: SNo w.
Assume Hwu1b: SNoLev w SNoLev u1.
Assume Hwu1c: w < u1.
We prove the intermediate claim LwSy: w SNoS_ (SNoLev y).
Apply SNoS_I2 w y Hwu1a Hy to the current goal.
We will prove SNoLev w SNoLev y.
An exact proof term for the current goal is ordinal_TransSet (SNoLev y) (SNoLev_ordinal y Hy) (SNoLev u1) Hu1b (SNoLev w) Hwu1b.
We prove the intermediate claim Lzw: SNo (z * w).
An exact proof term for the current goal is LRx z Hz w Hwu1a.
We prove the intermediate claim Lxw: SNo (x * w).
Apply IHy w LwSy to the current goal.
Assume H1 _ _ _ _.
An exact proof term for the current goal is H1.
We prove the intermediate claim L3a: z * v1 + x * w < x * v1 + z * w.
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume _ _ IHy2 _ _.
An exact proof term for the current goal is IHy2 z Hz w Hwv1.
We prove the intermediate claim L3b: x * u1 + z * w < z * u1 + x * w.
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ _ IHy4.
An exact proof term for the current goal is IHy4 z Hz w Hwu1.
An exact proof term for the current goal is add_SNo_Lt_subprop2 (x * u1) (z * v1) (z * u1) (x * v1) (z * w) (x * w) Lxu1 Lzv1 Lzu1 Lxv1 Lzw Lxw L3b L3a.
Assume Hv1u1: v1 SNoL u1.
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1) to the current goal.
Assume _ _ _ _ IHy4.
An exact proof term for the current goal is IHy4 z Hz v1 Hv1u1.
Assume Hu1v1: u1 SNoR v1.
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1) to the current goal.
Assume _ _ IHy2 _ _.
rewrite the current goal using add_SNo_com (x * u1) (z * v1) Lxu1 Lzv1 (from left to right).
rewrite the current goal using add_SNo_com (z * u1) (x * v1) Lzu1 Lxv1 (from left to right).
We will prove z * v1 + x * u1 < x * v1 + z * u1.
An exact proof term for the current goal is IHy2 z Hz u1 Hu1v1.
We prove the intermediate claim L3u0: x * u1 + u0 * v1 < u0 * u1 + x * v1.
An exact proof term for the current goal is L3 u0 Hu0.
We prove the intermediate claim L3v0: x * u1 + v0 * v1 < v0 * u1 + x * v1.
An exact proof term for the current goal is L3 v0 Hv0.
We prove the intermediate claim L3u0imp: u0 * y + v0 * v1 < v0 * y + u0 * v1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
Assume H1.
Apply add_SNo_Lt_subprop3a (u0 * y) (x * u1) (v0 * v1) (v0 * y) (x * v1 + u0 * u1) (u0 * v1) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) Lu0v1 H1 to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0.
We prove the intermediate claim L3v0imp: u0 * y + v0 * u1 < v0 * y + u0 * u1u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1.
Assume H1.
Apply add_SNo_Lt_subprop3b (u0 * y) (x * u1 + v0 * v1) (v0 * y) (x * v1) (u0 * u1) (v0 * u1) Lu0y (SNo_add_SNo (x * u1) (v0 * v1) Lxu1 Lv0v1) Lv0y Lxv1 Lu0u1 Lv0u1 to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
An exact proof term for the current goal is H1.
We will prove x * u1 + v0 * v1 < v0 * u1 + x * v1.
An exact proof term for the current goal is L3v0.
Apply SNoL_or_SNoR_impred v0 u0 Hv0a Hu0a to the current goal.
Assume Hv0u0: v0 = u0.
rewrite the current goal using Hv0u0 (from left to right).
We will prove u0 * y + (x * u1 + u0 * v1) < u0 * y + (x * v1 + u0 * u1).
Apply add_SNo_Lt2 (u0 * y) (x * u1 + u0 * v1) (x * v1 + u0 * u1) Lu0y (SNo_add_SNo (x * u1) (u0 * v1) Lxu1 Lu0v1) (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0.
Let z be given.
Assume Hzu0: z SNoL u0.
Assume Hzv0: z SNoR v0.
Apply SNoR_E v0 Hv0a z Hzv0 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev v0.
Assume Hzc: v0 < z.
We prove the intermediate claim Lz: z SNoR x.
Apply SNoR_I x Hx z Hza to the current goal.
We will prove SNoLev z SNoLev x.
An exact proof term for the current goal is ordinal_TransSet (SNoLev x) (SNoLev_ordinal x Hx) (SNoLev v0) Hv0b (SNoLev z) Hzb.
We will prove x < z.
An exact proof term for the current goal is SNoLt_tra x v0 z Hx Hv0a Hza Hv0c Hzc.
We prove the intermediate claim L4: u0 * y + z * u1 < z * y + u0 * u1.
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ IHx3 _.
An exact proof term for the current goal is IHx3 z Hzu0 u1 Hu1.
We prove the intermediate claim L5: z * y + v0 * u1 < v0 * y + z * u1.
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ _ IHx2 _ _.
An exact proof term for the current goal is IHx2 z Hzv0 u1 Hu1.
An exact proof term for the current goal is add_SNo_Lt_subprop3d (u0 * y) (x * u1 + v0 * v1) (v0 * y) (x * v1) (u0 * u1) (z * u1) (z * y) (v0 * u1) Lu0y (SNo_add_SNo (x * u1) (v0 * v1) Lxu1 Lv0v1) Lv0y Lxv1 Lu0u1 (LRx z Lz u1 Hu1a) (LRxy z Lz) Lv0u1 L4 L3v0 L5.
Assume Hv0u0: v0 SNoL u0.
Apply L3v0imp to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ IHx3 _.
An exact proof term for the current goal is IHx3 v0 Hv0u0 u1 Hu1.
Assume Hu0v0: u0 SNoR v0.
Apply L3v0imp to the current goal.
We will prove u0 * y + v0 * u1 < v0 * y + u0 * u1.
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ _ IHx2 _ _.
An exact proof term for the current goal is IHx2 u0 Hu0v0 u1 Hu1.
Let z be given.
Assume Hzu0: z SNoR u0.
Assume Hzv0: z SNoL v0.
Apply SNoR_E u0 Hu0a z Hzu0 to the current goal.
Assume Hza: SNo z.
Assume Hzb: SNoLev z SNoLev u0.
Assume Hzc: u0 < z.
We prove the intermediate claim Lz: z SNoR x.
Apply SNoR_I x Hx z Hza to the current goal.
We will prove SNoLev z SNoLev x.
An exact proof term for the current goal is ordinal_TransSet (SNoLev x) (SNoLev_ordinal x Hx) (SNoLev u0) Hu0b (SNoLev z) Hzb.
We will prove x < z.
An exact proof term for the current goal is SNoLt_tra x u0 z Hx Hu0a Hza Hu0c Hzc.
We prove the intermediate claim L6: u0 * y + z * v1 < z * y + u0 * v1.
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ _ IHx4.
An exact proof term for the current goal is IHx4 z Hzu0 v1 Hv1.
We prove the intermediate claim L7: z * y + v0 * v1 < v0 * y + z * v1.
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ IHx1 _ _ _.
An exact proof term for the current goal is IHx1 z Hzv0 v1 Hv1.
Apply add_SNo_Lt_subprop3c (u0 * y) (x * u1) (v0 * v1) (v0 * y) (x * v1 + u0 * u1) (z * v1) (z * y) (u0 * v1) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1) (u0 * u1) Lxv1 Lu0u1) (LRx z Lz v1 Hv1a) (LRxy z Lz) Lu0v1 L6 to the current goal.
We will prove x * u1 + u0 * v1 < x * v1 + u0 * u1.
rewrite the current goal using add_SNo_com (x * v1) (u0 * u1) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0.
We will prove z * y + v0 * v1 < v0 * y + z * v1.
An exact proof term for the current goal is L7.
Assume Hv0u0: v0 SNoR u0.
Apply L3u0imp to the current goal.
We will prove u0 * y + v0 * v1 < v0 * y + u0 * v1.
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0) y Hy to the current goal.
Assume _ _ _ _ IHx4.
An exact proof term for the current goal is IHx4 v0 Hv0u0 v1 Hv1.
Assume Hu0v0: u0 SNoL v0.
Apply L3u0imp to the current goal.
We will prove u0 * y + v0 * v1 < v0 * y + u0 * v1.
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0) y Hy to the current goal.
Assume _ IHx1 _ _ _.
An exact proof term for the current goal is IHx1 u0 Hu0v0 v1 Hv1.
We prove the intermediate claim Lxy: SNo (x * y).
rewrite the current goal using Hxy (from left to right).
We will prove SNo (SNoCut L R).
An exact proof term for the current goal is SNoCutP_SNo_SNoCut L R LLR1.
Let p be given.
Assume Hp.
We will prove p.
Apply Hp to the current goal.
An exact proof term for the current goal is Lxy.
We will prove uSNoL x, vSNoL y, u * y + x * v < x * y + u * v.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
We will prove u * y + x * v < x * y + u * v.
We prove the intermediate claim L1: u * y + x * v + - u * v < x * y.
rewrite the current goal using Hxy (from left to right).
We will prove u * y + x * v + - u * v < SNoCut L R.
Apply SNoCutP_SNoCut_L L R LLR1 to the current goal.
We will prove u * y + x * v + - u * v L.
An exact proof term for the current goal is HLI1 u Hu v Hv.
Apply add_SNo_minus_Lt1 (u * y + x * v) (u * v) (x * y) to the current goal.
An exact proof term for the current goal is SNo_add_SNo (u * y) (x * v) (LLx u Hu y Hy) (LxLy v Hv).
An exact proof term for the current goal is LLx u Hu v Hva.
An exact proof term for the current goal is Lxy.
We will prove (u * y + x * v) + - u * v < x * y.
rewrite the current goal using add_SNo_assoc (u * y) (x * v) (- (u * v)) (LLx u Hu y Hy) (LxLy v Hv) (SNo_minus_SNo (u * v) (LLx u Hu v Hva)) (from right to left).
An exact proof term for the current goal is L1.
We will prove uSNoR x, vSNoR y, u * y + x * v < x * y + u * v.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
We will prove u * y + x * v < x * y + u * v.
We prove the intermediate claim L1: u * y + x * v + - u * v < x * y.
rewrite the current goal using Hxy (from left to right).
We will prove u * y + x * v + - u * v < SNoCut L R.
Apply SNoCutP_SNoCut_L L R LLR1 to the current goal.
We will prove u * y + x * v + - u * v L.
An exact proof term for the current goal is HLI2 u Hu v Hv.
Apply add_SNo_minus_Lt1 (u * y + x * v) (u * v) (x * y) to the current goal.
An exact proof term for the current goal is SNo_add_SNo (u * y) (x * v) (LRx u Hu y Hy) (LxRy v Hv).
An exact proof term for the current goal is LRx u Hu v Hva.
An exact proof term for the current goal is Lxy.
We will prove (u * y + x * v) + - u * v < x * y.
rewrite the current goal using add_SNo_assoc (u * y) (x * v) (- (u * v)) (LRx u Hu y Hy) (LxRy v Hv) (SNo_minus_SNo (u * v) (LRx u Hu v Hva)) (from right to left).
An exact proof term for the current goal is L1.
We will prove uSNoL x, vSNoR y, x * y + u * v < u * y + x * v.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
We will prove x * y + u * v < u * y + x * v.
We prove the intermediate claim L1: x * y < u * y + x * v + - u * v.
rewrite the current goal using Hxy (from left to right).
We will prove SNoCut L R < u * y + x * v + - u * v.
Apply SNoCutP_SNoCut_R L R LLR1 to the current goal.
We will prove u * y + x * v + - u * v R.
An exact proof term for the current goal is HRI1 u Hu v Hv.
Apply add_SNo_minus_Lt2 (u * y + x * v) (u * v) (x * y) to the current goal.
An exact proof term for the current goal is SNo_add_SNo (u * y) (x * v) (LLx u Hu y Hy) (LxRy v Hv).
An exact proof term for the current goal is LLx u Hu v Hva.
An exact proof term for the current goal is Lxy.
We will prove x * y < (u * y + x * v) + - u * v.
rewrite the current goal using add_SNo_assoc (u * y) (x * v) (- (u * v)) (LLx u Hu y Hy) (LxRy v Hv) (SNo_minus_SNo (u * v) (LLx u Hu v Hva)) (from right to left).
An exact proof term for the current goal is L1.
We will prove uSNoR x, vSNoL y, x * y + u * v < u * y + x * v.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
We will prove x * y + u * v < u * y + x * v.
We prove the intermediate claim L1: x * y < u * y + x * v + - u * v.
rewrite the current goal using Hxy (from left to right).
We will prove SNoCut L R < u * y + x * v + - u * v.
Apply SNoCutP_SNoCut_R L R LLR1 to the current goal.
We will prove u * y + x * v + - u * v R.
An exact proof term for the current goal is HRI2 u Hu v Hv.
Apply add_SNo_minus_Lt2 (u * y + x * v) (u * v) (x * y) to the current goal.
An exact proof term for the current goal is SNo_add_SNo (u * y) (x * v) (LRx u Hu y Hy) (LxLy v Hv).
An exact proof term for the current goal is LRx u Hu v Hva.
An exact proof term for the current goal is Lxy.
We will prove x * y < (u * y + x * v) + - u * v.
rewrite the current goal using add_SNo_assoc (u * y) (x * v) (- (u * v)) (LRx u Hu y Hy) (LxLy v Hv) (SNo_minus_SNo (u * v) (LRx u Hu v Hva)) (from right to left).
An exact proof term for the current goal is L1.
Theorem. (SNo_mul_SNo) The following is provable:
∀x y, SNo xSNo ySNo (x * y)
Proof:
Let x and y be given.
Assume Hx Hy.
Apply mul_SNo_prop_1 x Hx y Hy to the current goal.
Assume H _ _ _ _.
An exact proof term for the current goal is H.
Theorem. (SNo_mul_SNo_lem) The following is provable:
∀x y u v, SNo xSNo ySNo uSNo vSNo (u * y + x * v + - (u * v))
Proof:
Let x, y, u and v be given.
Assume Hx Hy Hu Hv.
Apply SNo_add_SNo_3c to the current goal.
Apply SNo_mul_SNo to the current goal.
An exact proof term for the current goal is Hu.
An exact proof term for the current goal is Hy.
Apply SNo_mul_SNo to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hv.
Apply SNo_mul_SNo to the current goal.
An exact proof term for the current goal is Hu.
An exact proof term for the current goal is Hv.
Theorem. (SNo_mul_SNo_3) The following is provable:
∀x y z, SNo xSNo ySNo zSNo (x * y * z)
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
Apply SNo_mul_SNo to the current goal.
An exact proof term for the current goal is Hx.
Apply SNo_mul_SNo to the current goal.
An exact proof term for the current goal is Hy.
An exact proof term for the current goal is Hz.
Theorem. (mul_SNo_eq_3) The following is provable:
∀x y, SNo xSNo y∀p : prop, (∀L R, SNoCutP L R(∀u, u L(∀q : prop, (w0SNoL x, w1SNoL y, u = w0 * y + x * w1 + - w0 * w1q)(z0SNoR x, z1SNoR y, u = z0 * y + x * z1 + - z0 * z1q)q))(w0SNoL x, w1SNoL y, w0 * y + x * w1 + - w0 * w1 L)(z0SNoR x, z1SNoR y, z0 * y + x * z1 + - z0 * z1 L)(∀u, u R(∀q : prop, (w0SNoL x, z1SNoR y, u = w0 * y + x * z1 + - w0 * z1q)(z0SNoR x, w1SNoL y, u = z0 * y + x * w1 + - z0 * w1q)q))(w0SNoL x, z1SNoR y, w0 * y + x * z1 + - w0 * z1 R)(z0SNoR x, w1SNoL y, z0 * y + x * w1 + - z0 * w1 R)x * y = SNoCut L Rp)p
Proof:
Let x and y be given.
Assume Hx Hy.
Let p be given.
Assume Hp.
Apply mul_SNo_eq_2 x y Hx Hy to the current goal.
Let L and R be given.
Assume HLE HLI1 HLI2 HRE HRI1 HRI2 He.
Apply mul_SNo_prop_1 x Hx y Hy to the current goal.
Assume Hxy: SNo (x * y).
Assume Hxy1: uSNoL x, vSNoL y, u * y + x * v < x * y + u * v.
Assume Hxy2: uSNoR x, vSNoR y, u * y + x * v < x * y + u * v.
Assume Hxy3: uSNoL x, vSNoR y, x * y + u * v < u * y + x * v.
Assume Hxy4: uSNoR x, vSNoL y, x * y + u * v < u * y + x * v.
We prove the intermediate claim LLR: SNoCutP L R.
We will prove (wL, SNo w) (zR, SNo z) (wL, zR, w < z).
Apply and3I to the current goal.
Let w be given.
Assume Hw.
We will prove SNo w.
Apply HLE w Hw to the current goal.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv Hwe.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using Hwe (from left to right).
We will prove SNo (u * y + x * v + - u * v).
Apply SNo_mul_SNo_lem to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hy.
An exact proof term for the current goal is Hua.
An exact proof term for the current goal is Hva.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv Hwe.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using Hwe (from left to right).
We will prove SNo (u * y + x * v + - u * v).
Apply SNo_mul_SNo_lem to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hy.
An exact proof term for the current goal is Hua.
An exact proof term for the current goal is Hva.
Let z be given.
Assume Hz.
We will prove SNo z.
Apply HRE z Hz to the current goal.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv Hze.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using Hze (from left to right).
We will prove SNo (u * y + x * v + - u * v).
Apply SNo_mul_SNo_lem to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hy.
An exact proof term for the current goal is Hua.
An exact proof term for the current goal is Hva.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv Hze.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using Hze (from left to right).
We will prove SNo (u * y + x * v + - u * v).
Apply SNo_mul_SNo_lem to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hy.
An exact proof term for the current goal is Hua.
An exact proof term for the current goal is Hva.
Let w be given.
Assume Hw.
Let z be given.
Assume Hz.
We will prove w < z.
Apply HLE w Hw to the current goal.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv Hwe.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua: SNo u.
Assume _ _.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva: SNo v.
Assume _ _.
We prove the intermediate claim Luy: SNo (u * y).
An exact proof term for the current goal is SNo_mul_SNo u y Hua Hy.
We prove the intermediate claim Lxv: SNo (x * v).
An exact proof term for the current goal is SNo_mul_SNo x v Hx Hva.
We prove the intermediate claim Luv: SNo (u * v).
An exact proof term for the current goal is SNo_mul_SNo u v Hua Hva.
Apply HRE z Hz to the current goal.
Let u' be given.
Assume Hu'.
Let v' be given.
Assume Hv' Hze.
Apply SNoL_E x Hx u' Hu' to the current goal.
Assume Hu'a: SNo u'.
Assume _ _.
Apply SNoR_E y Hy v' Hv' to the current goal.
Assume Hv'a: SNo v'.
Assume _ _.
We prove the intermediate claim Lu'y: SNo (u' * y).
An exact proof term for the current goal is SNo_mul_SNo u' y Hu'a Hy.
We prove the intermediate claim Lxv': SNo (x * v').
An exact proof term for the current goal is SNo_mul_SNo x v' Hx Hv'a.
We prove the intermediate claim Lu'v': SNo (u' * v').
An exact proof term for the current goal is SNo_mul_SNo u' v' Hu'a Hv'a.
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v'.
Apply add_SNo_minus_Lt_lem (u * y) (x * v) (u * v) (u' * y) (x * v') (u' * v') Luy Lxv Luv Lu'y Lxv' Lu'v' to the current goal.
We will prove u * y + x * v + u' * v' < u' * y + x * v' + u * v.
Apply SNoLt_tra (u * y + x * v + u' * v') (x * y + u * v + u' * v') (u' * y + x * v' + u * v) (SNo_add_SNo_3 (u * y) (x * v) (u' * v') Luy Lxv Lu'v') (SNo_add_SNo_3 (x * y) (u * v) (u' * v') Hxy Luv Lu'v') (SNo_add_SNo_3 (u' * y) (x * v') (u * v) Lu'y Lxv' Luv) to the current goal.
We will prove u * y + x * v + u' * v' < x * y + u * v + u' * v'.
Apply add_SNo_3_3_3_Lt1 (u * y) (x * v) (x * y) (u * v) (u' * v') Luy Lxv Hxy Luv Lu'v' to the current goal.
We will prove u * y + x * v < x * y + u * v.
An exact proof term for the current goal is Hxy1 u Hu v Hv.
We will prove x * y + u * v + u' * v' < u' * y + x * v' + u * v.
Apply add_SNo_3_2_3_Lt1 (x * y) (u' * v') (u' * y) (x * v') (u * v) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will prove u' * v' + x * y < u' * y + x * v'.
rewrite the current goal using add_SNo_com (u' * v') (x * y) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy3 u' Hu' v' Hv'.
Let u' be given.
Assume Hu'.
Let v' be given.
Assume Hv' Hze.
Apply SNoR_E x Hx u' Hu' to the current goal.
Assume Hu'a: SNo u'.
Assume _ _.
Apply SNoL_E y Hy v' Hv' to the current goal.
Assume Hv'a: SNo v'.
Assume _ _.
We prove the intermediate claim Lu'y: SNo (u' * y).
An exact proof term for the current goal is SNo_mul_SNo u' y Hu'a Hy.
We prove the intermediate claim Lxv': SNo (x * v').
An exact proof term for the current goal is SNo_mul_SNo x v' Hx Hv'a.
We prove the intermediate claim Lu'v': SNo (u' * v').
An exact proof term for the current goal is SNo_mul_SNo u' v' Hu'a Hv'a.
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v'.
Apply add_SNo_minus_Lt_lem (u * y) (x * v) (u * v) (u' * y) (x * v') (u' * v') Luy Lxv Luv Lu'y Lxv' Lu'v' to the current goal.
We will prove u * y + x * v + u' * v' < u' * y + x * v' + u * v.
Apply SNoLt_tra (u * y + x * v + u' * v') (x * y + u * v + u' * v') (u' * y + x * v' + u * v) (SNo_add_SNo_3 (u * y) (x * v) (u' * v') Luy Lxv Lu'v') (SNo_add_SNo_3 (x * y) (u * v) (u' * v') Hxy Luv Lu'v') (SNo_add_SNo_3 (u' * y) (x * v') (u * v) Lu'y Lxv' Luv) to the current goal.
We will prove u * y + x * v + u' * v' < x * y + u * v + u' * v'.
Apply add_SNo_3_3_3_Lt1 (u * y) (x * v) (x * y) (u * v) (u' * v') Luy Lxv Hxy Luv Lu'v' to the current goal.
We will prove u * y + x * v < x * y + u * v.
An exact proof term for the current goal is Hxy1 u Hu v Hv.
We will prove x * y + u * v + u' * v' < u' * y + x * v' + u * v.
Apply add_SNo_3_2_3_Lt1 (x * y) (u' * v') (u' * y) (x * v') (u * v) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will prove u' * v' + x * y < u' * y + x * v'.
rewrite the current goal using add_SNo_com (u' * v') (x * y) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy4 u' Hu' v' Hv'.
Let u be given.
Assume Hu.
Let v be given.
Assume Hv Hwe.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua: SNo u.
Assume _ _.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva: SNo v.
Assume _ _.
We prove the intermediate claim Luy: SNo (u * y).
An exact proof term for the current goal is SNo_mul_SNo u y Hua Hy.
We prove the intermediate claim Lxv: SNo (x * v).
An exact proof term for the current goal is SNo_mul_SNo x v Hx Hva.
We prove the intermediate claim Luv: SNo (u * v).
An exact proof term for the current goal is SNo_mul_SNo u v Hua Hva.
Apply HRE z Hz to the current goal.
Let u' be given.
Assume Hu'.
Let v' be given.
Assume Hv' Hze.
Apply SNoL_E x Hx u' Hu' to the current goal.
Assume Hu'a: SNo u'.
Assume _ _.
Apply SNoR_E y Hy v' Hv' to the current goal.
Assume Hv'a: SNo v'.
Assume _ _.
We prove the intermediate claim Lu'y: SNo (u' * y).
An exact proof term for the current goal is SNo_mul_SNo u' y Hu'a Hy.
We prove the intermediate claim Lxv': SNo (x * v').
An exact proof term for the current goal is SNo_mul_SNo x v' Hx Hv'a.
We prove the intermediate claim Lu'v': SNo (u' * v').
An exact proof term for the current goal is SNo_mul_SNo u' v' Hu'a Hv'a.
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v'.
Apply add_SNo_minus_Lt_lem (u * y) (x * v) (u * v) (u' * y) (x * v') (u' * v') Luy Lxv Luv Lu'y Lxv' Lu'v' to the current goal.
We will prove u * y + x * v + u' * v' < u' * y + x * v' + u * v.
Apply SNoLt_tra (u * y + x * v + u' * v') (x * y + u * v + u' * v') (u' * y + x * v' + u * v) (SNo_add_SNo_3 (u * y) (x * v) (u' * v') Luy Lxv Lu'v') (SNo_add_SNo_3 (x * y) (u * v) (u' * v') Hxy Luv Lu'v') (SNo_add_SNo_3 (u' * y) (x * v') (u * v) Lu'y Lxv' Luv) to the current goal.
We will prove u * y + x * v + u' * v' < x * y + u * v + u' * v'.
Apply add_SNo_3_3_3_Lt1 (u * y) (x * v) (x * y) (u * v) (u' * v') Luy Lxv Hxy Luv Lu'v' to the current goal.
We will prove u * y + x * v < x * y + u * v.
An exact proof term for the current goal is Hxy2 u Hu v Hv.
We will prove x * y + u * v + u' * v' < u' * y + x * v' + u * v.
Apply add_SNo_3_2_3_Lt1 (x * y) (u' * v') (u' * y) (x * v') (u * v) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will prove u' * v' + x * y < u' * y + x * v'.
rewrite the current goal using add_SNo_com (u' * v') (x * y) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy3 u' Hu' v' Hv'.
Let u' be given.
Assume Hu'.
Let v' be given.
Assume Hv' Hze.
Apply SNoR_E x Hx u' Hu' to the current goal.
Assume Hu'a: SNo u'.
Assume _ _.
Apply SNoL_E y Hy v' Hv' to the current goal.
Assume Hv'a: SNo v'.
Assume _ _.
We prove the intermediate claim Lu'y: SNo (u' * y).
An exact proof term for the current goal is SNo_mul_SNo u' y Hu'a Hy.
We prove the intermediate claim Lxv': SNo (x * v').
An exact proof term for the current goal is SNo_mul_SNo x v' Hx Hv'a.
We prove the intermediate claim Lu'v': SNo (u' * v').
An exact proof term for the current goal is SNo_mul_SNo u' v' Hu'a Hv'a.
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v'.
Apply add_SNo_minus_Lt_lem (u * y) (x * v) (u * v) (u' * y) (x * v') (u' * v') Luy Lxv Luv Lu'y Lxv' Lu'v' to the current goal.
We will prove u * y + x * v + u' * v' < u' * y + x * v' + u * v.
Apply SNoLt_tra (u * y + x * v + u' * v') (x * y + u * v + u' * v') (u' * y + x * v' + u * v) (SNo_add_SNo_3 (u * y) (x * v) (u' * v') Luy Lxv Lu'v') (SNo_add_SNo_3 (x * y) (u * v) (u' * v') Hxy Luv Lu'v') (SNo_add_SNo_3 (u' * y) (x * v') (u * v) Lu'y Lxv' Luv) to the current goal.
We will prove u * y + x * v + u' * v' < x * y + u * v + u' * v'.
Apply add_SNo_3_3_3_Lt1 (u * y) (x * v) (x * y) (u * v) (u' * v') Luy Lxv Hxy Luv Lu'v' to the current goal.
We will prove u * y + x * v < x * y + u * v.
An exact proof term for the current goal is Hxy2 u Hu v Hv.
We will prove x * y + u * v + u' * v' < u' * y + x * v' + u * v.
Apply add_SNo_3_2_3_Lt1 (x * y) (u' * v') (u' * y) (x * v') (u * v) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will prove u' * v' + x * y < u' * y + x * v'.
rewrite the current goal using add_SNo_com (u' * v') (x * y) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy4 u' Hu' v' Hv'.
An exact proof term for the current goal is Hp L R LLR HLE HLI1 HLI2 HRE HRI1 HRI2 He.
Theorem. (mul_SNo_Lt) The following is provable:
∀x y u v, SNo xSNo ySNo uSNo vu < xv < yu * y + x * v < x * y + u * v
Proof:
Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy.
Apply mul_SNo_prop_1 x Hx y Hy to the current goal.
Assume Hxy: SNo (x * y).
Assume Hxy1: uSNoL x, vSNoL y, u * y + x * v < x * y + u * v.
Assume Hxy2: uSNoR x, vSNoR y, u * y + x * v < x * y + u * v.
Assume Hxy3: uSNoL x, vSNoR y, x * y + u * v < u * y + x * v.
Assume Hxy4: uSNoR x, vSNoL y, x * y + u * v < u * y + x * v.
Apply mul_SNo_prop_1 u Hu y Hy to the current goal.
Assume Huy: SNo (u * y).
Assume Huy1: xSNoL u, vSNoL y, x * y + u * v < u * y + x * v.
Assume Huy2: xSNoR u, vSNoR y, x * y + u * v < u * y + x * v.
Assume Huy3: xSNoL u, vSNoR y, u * y + x * v < x * y + u * v.
Assume Huy4: xSNoR u, vSNoL y, u * y + x * v < x * y + u * v.
Apply mul_SNo_prop_1 x Hx v Hv to the current goal.
Assume Hxv: SNo (x * v).
Assume Hxv1: uSNoL x, ySNoL v, u * v + x * y < x * v + u * y.
Assume Hxv2: uSNoR x, ySNoR v, u * v + x * y < x * v + u * y.
Assume Hxv3: uSNoL x, ySNoR v, x * v + u * y < u * v + x * y.
Assume Hxv4: uSNoR x, ySNoL v, x * v + u * y < u * v + x * y.
Apply mul_SNo_prop_1 u Hu v Hv to the current goal.
Assume Huv: SNo (u * v).
Assume Huv1: xSNoL u, ySNoL v, x * v + u * y < u * v + x * y.
Assume Huv2: xSNoR u, ySNoR v, x * v + u * y < u * v + x * y.
Assume Huv3: xSNoL u, ySNoR v, u * v + x * y < x * v + u * y.
Assume Huv4: xSNoR u, ySNoL v, u * v + x * y < x * v + u * y.
We prove the intermediate claim Luyxv: SNo (u * y + x * v).
An exact proof term for the current goal is SNo_add_SNo (u * y) (x * v) Huy Hxv.
We prove the intermediate claim Lxyuv: SNo (x * y + u * v).
An exact proof term for the current goal is SNo_add_SNo (x * y) (u * v) Hxy Huv.
Apply SNoLtE u x Hu Hx Hux to the current goal.
Let w be given.
Assume Hw1: SNo w.
Assume Hw2: SNoLev w SNoLev u SNoLev x.
Assume Hw3: SNoEq_ (SNoLev w) w u.
Assume Hw4: SNoEq_ (SNoLev w) w x.
Assume Hw5: u < w.
Assume Hw6: w < x.
Assume Hw7: SNoLev w u.
Assume Hw8: SNoLev w x.
Apply binintersectE (SNoLev u) (SNoLev x) (SNoLev w) Hw2 to the current goal.
Assume Hw2a Hw2b.
We prove the intermediate claim Lwx: w SNoL x.
An exact proof term for the current goal is SNoL_I x Hx w Hw1 Hw2b Hw6.
We prove the intermediate claim Lwu: w SNoR u.
An exact proof term for the current goal is SNoR_I u Hu w Hw1 Hw2a Hw5.
We prove the intermediate claim Lwy: SNo (w * y).
An exact proof term for the current goal is SNo_mul_SNo w y Hw1 Hy.
We prove the intermediate claim Lwv: SNo (w * v).
An exact proof term for the current goal is SNo_mul_SNo w v Hw1 Hv.
We prove the intermediate claim Lwyxv: SNo (w * y + x * v).
An exact proof term for the current goal is SNo_add_SNo (w * y) (x * v) Lwy Hxv.
We prove the intermediate claim Lwyuv: SNo (w * y + u * v).
An exact proof term for the current goal is SNo_add_SNo (w * y) (u * v) Lwy Huv.
We prove the intermediate claim Lxywv: SNo (x * y + w * v).
An exact proof term for the current goal is SNo_add_SNo (x * y) (w * v) Hxy Lwv.
We prove the intermediate claim Luywv: SNo (u * y + w * v).
An exact proof term for the current goal is SNo_add_SNo (u * y) (w * v) Huy Lwv.
We prove the intermediate claim Luvwy: SNo (u * v + w * y).
An exact proof term for the current goal is SNo_add_SNo (u * v) (w * y) Huv Lwy.
We prove the intermediate claim Lwvxy: SNo (w * v + x * y).
An exact proof term for the current goal is SNo_add_SNo (w * v) (x * y) Lwv Hxy.
We prove the intermediate claim Lxvwy: SNo (x * v + w * y).
An exact proof term for the current goal is SNo_add_SNo (x * v) (w * y) Hxv Lwy.
We prove the intermediate claim Lwvuy: SNo (w * v + u * y).
An exact proof term for the current goal is SNo_add_SNo (w * v) (u * y) Lwv Huy.
Apply SNoLtE v y Hv Hy Hvy to the current goal.
Let z be given.
Assume Hz1: SNo z.
Assume Hz2: SNoLev z SNoLev v SNoLev y.
Assume Hz3: SNoEq_ (SNoLev z) z v.
Assume Hz4: SNoEq_ (SNoLev z) z y.
Assume Hz5: v < z.
Assume Hz6: z < y.
Assume Hz7: SNoLev z v.
Assume Hz8: SNoLev z y.
Apply binintersectE (SNoLev v) (SNoLev y) (SNoLev z) Hz2 to the current goal.
Assume Hz2a Hz2b.
We prove the intermediate claim Lzy: z SNoL y.
An exact proof term for the current goal is SNoL_I y Hy z Hz1 Hz2b Hz6.
We prove the intermediate claim Lzv: z SNoR v.
An exact proof term for the current goal is SNoR_I v Hv z Hz1 Hz2a Hz5.
We prove the intermediate claim Lxz: SNo (x * z).
An exact proof term for the current goal is SNo_mul_SNo x z Hx Hz1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu Hz1.
We prove the intermediate claim Lwz: SNo (w * z).
An exact proof term for the current goal is SNo_mul_SNo w z Hw1 Hz1.
We prove the intermediate claim L1: w * y + x * z < x * y + w * z.
An exact proof term for the current goal is Hxy1 w Lwx z Lzy.
We prove the intermediate claim L2: w * v + u * z < u * v + w * z.
An exact proof term for the current goal is Huv2 w Lwu z Lzv.
We prove the intermediate claim L3: x * v + w * z < w * v + x * z.
An exact proof term for the current goal is Hxv3 w Lwx z Lzv.
We prove the intermediate claim L4: u * y + w * z < w * y + u * z.
An exact proof term for the current goal is Huy4 w Lwu z Lzy.
We prove the intermediate claim Lwzwz: SNo (w * z + w * z).
An exact proof term for the current goal is SNo_add_SNo (w * z) (w * z) Lwz Lwz.
Apply add_SNo_Lt1_cancel (u * y + x * v) (w * z + w * z) (x * y + u * v) Luyxv Lwzwz Lxyuv to the current goal.
We will prove (u * y + x * v) + (w * z + w * z) < (x * y + u * v) + (w * z + w * z).
We prove the intermediate claim Lwyuz: SNo (w * y + u * z).
An exact proof term for the current goal is SNo_add_SNo (w * y) (u * z) Lwy Luz.
We prove the intermediate claim Lwvxz: SNo (w * v + x * z).
An exact proof term for the current goal is SNo_add_SNo (w * v) (x * z) Lwv Lxz.
We prove the intermediate claim Luywz: SNo (u * y + w * z).
An exact proof term for the current goal is SNo_add_SNo (u * y) (w * z) Huy Lwz.
We prove the intermediate claim Lxvwz: SNo (x * v + w * z).
An exact proof term for the current goal is SNo_add_SNo (x * v) (w * z) Hxv Lwz.
We prove the intermediate claim Lwvuz: SNo (w * v + u * z).
An exact proof term for the current goal is SNo_add_SNo (w * v) (u * z) Lwv Luz.
We prove the intermediate claim Lxyxz: SNo (x * y + x * z).
An exact proof term for the current goal is SNo_add_SNo (x * y) (x * z) Hxy Lxz.
We prove the intermediate claim Lwyxz: SNo (w * y + x * z).
An exact proof term for the current goal is SNo_add_SNo (w * y) (x * z) Lwy Lxz.
We prove the intermediate claim Lxywz: SNo (x * y + w * z).
An exact proof term for the current goal is SNo_add_SNo (x * y) (w * z) Hxy Lwz.
We prove the intermediate claim Luvwz: SNo (u * v + w * z).
An exact proof term for the current goal is SNo_add_SNo (u * v) (w * z) Huv Lwz.
Apply SNoLt_tra ((u * y + x * v) + (w * z + w * z)) ((w * y + u * z) + (w * v + x * z)) ((x * y + u * v) + (w * z + w * z)) (SNo_add_SNo (u * y + x * v) (w * z + w * z) Luyxv Lwzwz) (SNo_add_SNo (w * y + u * z) (w * v + x * z) Lwyuz Lwvxz) (SNo_add_SNo (x * y + u * v) (w * z + w * z) Lxyuv Lwzwz) to the current goal.
We will prove (u * y + x * v) + (w * z + w * z) < (w * y + u * z) + (w * v + x * z).
rewrite the current goal using add_SNo_com_4_inner_mid (u * y) (x * v) (w * z) (w * z) Huy Hxv Lwz Lwz (from left to right).
We will prove (u * y + w * z) + (x * v + w * z) < (w * y + u * z) + (w * v + x * z).
Apply add_SNo_Lt3 (u * y + w * z) (x * v + w * z) (w * y + u * z) (w * v + x * z) Luywz Lxvwz Lwyuz Lwvxz to the current goal.
We will prove u * y + w * z < w * y + u * z.
An exact proof term for the current goal is L4.
We will prove x * v + w * z < w * v + x * z.
An exact proof term for the current goal is L3.
We will prove (w * y + u * z) + (w * v + x * z) < (x * y + u * v) + (w * z + w * z).
rewrite the current goal using add_SNo_com_4_inner_mid (x * y) (u * v) (w * z) (w * z) Hxy Huv Lwz Lwz (from left to right).
We will prove (w * y + u * z) + (w * v + x * z) < (x * y + w * z) + (u * v + w * z).
rewrite the current goal using add_SNo_com (w * y) (u * z) Lwy Luz (from left to right).
rewrite the current goal using add_SNo_com_4_inner_mid (u * z) (w * y) (w * v) (x * z) Luz Lwy Lwv Lxz (from left to right).
rewrite the current goal using add_SNo_com (w * v) (u * z) Lwv Luz (from right to left).
We will prove (w * v + u * z) + (w * y + x * z) < (x * y + w * z) + (u * v + w * z).
rewrite the current goal using add_SNo_com (w * v + u * z) (w * y + x * z) Lwvuz Lwyxz (from left to right).
We will prove (w * y + x * z) + (w * v + u * z) < (x * y + w * z) + (u * v + w * z).
Apply add_SNo_Lt3 (w * y + x * z) (w * v + u * z) (x * y + w * z) (u * v + w * z) Lwyxz Lwvuz Lxywz Luvwz to the current goal.
An exact proof term for the current goal is L1.
An exact proof term for the current goal is L2.
Assume H4: SNoLev v SNoLev y.
Assume H5: SNoEq_ (SNoLev v) v y.
Assume H6: SNoLev v y.
We prove the intermediate claim Lvy: v SNoL y.
An exact proof term for the current goal is SNoL_I y Hy v Hv H4 Hvy.
We prove the intermediate claim L1: w * y + x * v < x * y + w * v.
An exact proof term for the current goal is Hxy1 w Lwx v Lvy.
We prove the intermediate claim L2: u * y + w * v < w * y + u * v.
An exact proof term for the current goal is Huy4 w Lwu v Lvy.
We will prove u * y + x * v < x * y + u * v.
Apply add_SNo_Lt2_cancel (w * y) (u * y + x * v) (x * y + u * v) Lwy Luyxv Lxyuv to the current goal.
We will prove w * y + u * y + x * v < w * y + x * y + u * v.
Apply SNoLt_tra (w * y + u * y + x * v) (u * y + w * v + x * y) (w * y + x * y + u * v) (SNo_add_SNo (w * y) (u * y + x * v) Lwy Luyxv) (SNo_add_SNo (u * y) (w * v + x * y) Huy (SNo_add_SNo (w * v) (x * y) Lwv Hxy)) (SNo_add_SNo (w * y) (x * y + u * v) Lwy Lxyuv) to the current goal.
We will prove w * y + u * y + x * v < u * y + w * v + x * y.
rewrite the current goal using add_SNo_com_3_0_1 (w * y) (u * y) (x * v) Lwy Huy Hxv (from left to right).
We will prove u * y + w * y + x * v < u * y + w * v + x * y.
rewrite the current goal using add_SNo_com (w * v) (x * y) Lwv Hxy (from left to right).
An exact proof term for the current goal is add_SNo_Lt2 (u * y) (w * y + x * v) (x * y + w * v) Huy Lwyxv Lxywv L1.
We will prove u * y + w * v + x * y < w * y + x * y + u * v.
rewrite the current goal using add_SNo_assoc (u * y) (w * v) (x * y) Huy Lwv Hxy (from left to right).
We will prove (u * y + w * v) + x * y < w * y + x * y + u * v.
rewrite the current goal using add_SNo_com (x * y) (u * v) Hxy Huv (from left to right).
We will prove (u * y + w * v) + x * y < w * y + u * v + x * y.
rewrite the current goal using add_SNo_assoc (w * y) (u * v) (x * y) Lwy Huv Hxy (from left to right).
We will prove (u * y + w * v) + x * y < (w * y + u * v) + x * y.
An exact proof term for the current goal is add_SNo_Lt1 (u * y + w * v) (x * y) (w * y + u * v) Luywv Hxy Lwyuv L2.
Assume H4: SNoLev y SNoLev v.
Assume H5: SNoEq_ (SNoLev y) v y.
Assume H6: SNoLev y v.
We prove the intermediate claim Lyv: y SNoR v.
An exact proof term for the current goal is SNoR_I v Hv y Hy H4 Hvy.
We prove the intermediate claim L1: x * v + w * y < w * v + x * y.
An exact proof term for the current goal is Hxv3 w Lwx y Lyv.
We prove the intermediate claim L2: w * v + u * y < u * v + w * y.
An exact proof term for the current goal is Huv2 w Lwu y Lyv.
We will prove u * y + x * v < x * y + u * v.
Apply add_SNo_Lt2_cancel (w * y) (u * y + x * v) (x * y + u * v) Lwy Luyxv Lxyuv to the current goal.
We will prove w * y + u * y + x * v < w * y + x * y + u * v.
Apply SNoLt_tra (w * y + u * y + x * v) (u * y + w * v + x * y) (w * y + x * y + u * v) (SNo_add_SNo (w * y) (u * y + x * v) Lwy Luyxv) (SNo_add_SNo (u * y) (w * v + x * y) Huy (SNo_add_SNo (w * v) (x * y) Lwv Hxy)) (SNo_add_SNo (w * y) (x * y + u * v) Lwy Lxyuv) to the current goal.
We will prove w * y + u * y + x * v < u * y + w * v + x * y.
rewrite the current goal using add_SNo_rotate_3_1 (u * y) (x * v) (w * y) Huy Hxv Lwy (from right to left).
We will prove u * y + x * v + w * y < u * y + w * v + x * y.
An exact proof term for the current goal is add_SNo_Lt2 (u * y) (x * v + w * y) (w * v + x * y) Huy Lxvwy Lwvxy L1.
We will prove u * y + w * v + x * y < w * y + x * y + u * v.
rewrite the current goal using add_SNo_rotate_3_1 (w * y) (x * y) (u * v) Lwy Hxy Huv (from left to right).
We will prove u * y + w * v + x * y < u * v + w * y + x * y.
rewrite the current goal using add_SNo_assoc (u * y) (w * v) (x * y) Huy Lwv Hxy (from left to right).
rewrite the current goal using add_SNo_assoc (u * v) (w * y) (x * y) Huv Lwy Hxy (from left to right).
We will prove (u * y + w * v) + x * y < (u * v + w * y) + x * y.
rewrite the current goal using add_SNo_com (u * y) (w * v) Huy Lwv (from left to right).
We will prove (w * v + u * y) + x * y < (u * v + w * y) + x * y.
Apply add_SNo_Lt1 (w * v + u * y) (x * y) (u * v + w * y) Lwvuy Hxy Luvwy L2 to the current goal.
Assume H1: SNoLev u SNoLev x.
Assume H2: SNoEq_ (SNoLev u) u x.
Assume H3: SNoLev u x.
We prove the intermediate claim Lux: u SNoL x.
An exact proof term for the current goal is SNoL_I x Hx u Hu H1 Hux.
Apply SNoLtE v y Hv Hy Hvy to the current goal.
Let z be given.
Assume Hz1: SNo z.
Assume Hz2: SNoLev z SNoLev v SNoLev y.
Assume Hz3: SNoEq_ (SNoLev z) z v.
Assume Hz4: SNoEq_ (SNoLev z) z y.
Assume Hz5: v < z.
Assume Hz6: z < y.
Assume Hz7: SNoLev z v.
Assume Hz8: SNoLev z y.
Apply binintersectE (SNoLev v) (SNoLev y) (SNoLev z) Hz2 to the current goal.
Assume Hz2a Hz2b.
We prove the intermediate claim Lzy: z SNoL y.
An exact proof term for the current goal is SNoL_I y Hy z Hz1 Hz2b Hz6.
We prove the intermediate claim Lzv: z SNoR v.
An exact proof term for the current goal is SNoR_I v Hv z Hz1 Hz2a Hz5.
We prove the intermediate claim Lxz: SNo (x * z).
An exact proof term for the current goal is SNo_mul_SNo x z Hx Hz1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu Hz1.
We prove the intermediate claim Luyxz: SNo (u * y + x * z).
An exact proof term for the current goal is SNo_add_SNo (u * y) (x * z) Huy Lxz.
We prove the intermediate claim Luvxz: SNo (u * v + x * z).
An exact proof term for the current goal is SNo_add_SNo (u * v) (x * z) Huv Lxz.
We prove the intermediate claim Lxyuz: SNo (x * y + u * z).
An exact proof term for the current goal is SNo_add_SNo (x * y) (u * z) Hxy Luz.
We prove the intermediate claim Lxvuz: SNo (x * v + u * z).
An exact proof term for the current goal is SNo_add_SNo (x * v) (u * z) Hxv Luz.
We prove the intermediate claim L1: u * y + x * z < x * y + u * z.
An exact proof term for the current goal is Hxy1 u Lux z Lzy.
We prove the intermediate claim L2: x * v + u * z < u * v + x * z.
An exact proof term for the current goal is Hxv3 u Lux z Lzv.
We will prove u * y + x * v < x * y + u * v.
Apply add_SNo_Lt2_cancel (x * z) (u * y + x * v) (x * y + u * v) Lxz Luyxv Lxyuv to the current goal.
We will prove x * z + u * y + x * v < x * z + x * y + u * v.
Apply SNoLt_tra (x * z + u * y + x * v) (x * y + u * z + x * v) (x * z + x * y + u * v) (SNo_add_SNo (x * z) (u * y + x * v) Lxz Luyxv) (SNo_add_SNo (x * y) (u * z + x * v) Hxy (SNo_add_SNo (u * z) (x * v) Luz Hxv)) (SNo_add_SNo (x * z) (x * y + u * v) Lxz Lxyuv) to the current goal.
We will prove x * z + u * y + x * v < x * y + u * z + x * v.
rewrite the current goal using add_SNo_assoc (x * z) (u * y) (x * v) Lxz Huy Hxv (from left to right).
rewrite the current goal using add_SNo_com (x * z) (u * y) Lxz Huy (from left to right).
We will prove (u * y + x * z) + x * v < x * y + u * z + x * v.
rewrite the current goal using add_SNo_assoc (x * y) (u * z) (x * v) Hxy Luz Hxv (from left to right).
An exact proof term for the current goal is add_SNo_Lt1 (u * y + x * z) (x * v) (x * y + u * z) Luyxz Hxv Lxyuz L1.
We will prove x * y + u * z + x * v < x * z + x * y + u * v.
rewrite the current goal using add_SNo_rotate_3_1 (x * y) (u * v) (x * z) Hxy Huv Lxz (from right to left).
We will prove x * y + u * z + x * v < x * y + u * v + x * z.
rewrite the current goal using add_SNo_com (u * z) (x * v) Luz Hxv (from left to right).
We will prove x * y + x * v + u * z < x * y + u * v + x * z.
An exact proof term for the current goal is add_SNo_Lt2 (x * y) (x * v + u * z) (u * v + x * z) Hxy Lxvuz Luvxz L2.
Assume H4: SNoLev v SNoLev y.
Assume H5: SNoEq_ (SNoLev v) v y.
Assume H6: SNoLev v y.
We prove the intermediate claim Lvy: v SNoL y.
An exact proof term for the current goal is SNoL_I y Hy v Hv H4 Hvy.
An exact proof term for the current goal is Hxy1 u Lux v Lvy.
Assume H4: SNoLev y SNoLev v.
Assume H5: SNoEq_ (SNoLev y) v y.
Assume H6: SNoLev y v.
We prove the intermediate claim Lyv: y SNoR v.
An exact proof term for the current goal is SNoR_I v Hv y Hy H4 Hvy.
We will prove u * y + x * v < x * y + u * v.
rewrite the current goal using add_SNo_com (u * y) (x * v) Huy Hxv (from left to right).
rewrite the current goal using add_SNo_com (x * y) (u * v) Hxy Huv (from left to right).
An exact proof term for the current goal is Hxv3 u Lux y Lyv.
Assume H1: SNoLev x SNoLev u.
Assume H2: SNoEq_ (SNoLev x) u x.
Assume H3: SNoLev x u.
We prove the intermediate claim Lxu: x SNoR u.
An exact proof term for the current goal is SNoR_I u Hu x Hx H1 Hux.
Apply SNoLtE v y Hv Hy Hvy to the current goal.
Let z be given.
Assume Hz1: SNo z.
Assume Hz2: SNoLev z SNoLev v SNoLev y.
Assume Hz3: SNoEq_ (SNoLev z) z v.
Assume Hz4: SNoEq_ (SNoLev z) z y.
Assume Hz5: v < z.
Assume Hz6: z < y.
Assume Hz7: SNoLev z v.
Assume Hz8: SNoLev z y.
Apply binintersectE (SNoLev v) (SNoLev y) (SNoLev z) Hz2 to the current goal.
Assume Hz2a Hz2b.
We prove the intermediate claim Lzy: z SNoL y.
An exact proof term for the current goal is SNoL_I y Hy z Hz1 Hz2b Hz6.
We prove the intermediate claim Lzv: z SNoR v.
An exact proof term for the current goal is SNoR_I v Hv z Hz1 Hz2a Hz5.
We prove the intermediate claim Lxz: SNo (x * z).
An exact proof term for the current goal is SNo_mul_SNo x z Hx Hz1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu Hz1.
We prove the intermediate claim L1: u * y + x * z < x * y + u * z.
An exact proof term for the current goal is Huy4 x Lxu z Lzy.
We prove the intermediate claim L2: x * v + u * z < u * v + x * z.
An exact proof term for the current goal is Huv2 x Lxu z Lzv.
We will prove u * y + x * v < x * y + u * v.
Apply add_SNo_Lt1_cancel (u * y + x * v) (x * z) (x * y + u * v) Luyxv Lxz Lxyuv to the current goal.
We will prove (u * y + x * v) + x * z < (x * y + u * v) + x * z.
rewrite the current goal using add_SNo_com (u * y) (x * v) Huy Hxv (from left to right).
We will prove (x * v + u * y) + x * z < (x * y + u * v) + x * z.
rewrite the current goal using add_SNo_assoc (x * v) (u * y) (x * z) Hxv Huy Lxz (from right to left).
rewrite the current goal using add_SNo_assoc (x * y) (u * v) (x * z) Hxy Huv Lxz (from right to left).
We will prove x * v + u * y + x * z < x * y + u * v + x * z.
We prove the intermediate claim Luyxz: SNo (u * y + x * z).
An exact proof term for the current goal is SNo_add_SNo (u * y) (x * z) Huy Lxz.
We prove the intermediate claim Luzxy: SNo (u * z + x * y).
An exact proof term for the current goal is SNo_add_SNo (u * z) (x * y) Luz Hxy.
We prove the intermediate claim Luvxz: SNo (u * v + x * z).
An exact proof term for the current goal is SNo_add_SNo (u * v) (x * z) Huv Lxz.
Apply SNoLt_tra (x * v + u * y + x * z) (x * v + u * z + x * y) (x * y + u * v + x * z) (SNo_add_SNo (x * v) (u * y + x * z) Hxv Luyxz) (SNo_add_SNo (x * v) (u * z + x * y) Hxv Luzxy) (SNo_add_SNo (x * y) (u * v + x * z) Hxy Luvxz) to the current goal.
We will prove x * v + u * y + x * z < x * v + u * z + x * y.
rewrite the current goal using add_SNo_com (u * z) (x * y) Luz Hxy (from left to right).
An exact proof term for the current goal is add_SNo_Lt2 (x * v) (u * y + x * z) (x * y + u * z) Hxv Luyxz (SNo_add_SNo (x * y) (u * z) Hxy Luz) L1.
We will prove x * v + u * z + x * y < x * y + u * v + x * z.
rewrite the current goal using add_SNo_rotate_3_1 (x * v) (u * z) (x * y) Hxv Luz Hxy (from left to right).
We will prove x * y + x * v + u * z < x * y + u * v + x * z.
An exact proof term for the current goal is add_SNo_Lt2 (x * y) (x * v + u * z) (u * v + x * z) Hxy (SNo_add_SNo (x * v) (u * z) Hxv Luz) Luvxz L2.
Assume H4: SNoLev v SNoLev y.
Assume H5: SNoEq_ (SNoLev v) v y.
Assume H6: SNoLev v y.
We prove the intermediate claim Lvy: v SNoL y.
An exact proof term for the current goal is SNoL_I y Hy v Hv H4 Hvy.
We will prove u * y + x * v < x * y + u * v.
An exact proof term for the current goal is Huy4 x Lxu v Lvy.
Assume H4: SNoLev y SNoLev v.
Assume H5: SNoEq_ (SNoLev y) v y.
Assume H6: SNoLev y v.
We prove the intermediate claim Lyv: y SNoR v.
An exact proof term for the current goal is SNoR_I v Hv y Hy H4 Hvy.
We will prove u * y + x * v < x * y + u * v.
rewrite the current goal using add_SNo_com (u * y) (x * v) Huy Hxv (from left to right).
rewrite the current goal using add_SNo_com (x * y) (u * v) Hxy Huv (from left to right).
An exact proof term for the current goal is Huv2 x Lxu y Lyv.
Theorem. (mul_SNo_Le) The following is provable:
∀x y u v, SNo xSNo ySNo uSNo vu xv yu * y + x * v x * y + u * v
Proof:
Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy.
Apply SNoLeE u x Hu Hx Hux to the current goal.
Assume Hux: u < x.
Apply SNoLeE v y Hv Hy Hvy to the current goal.
Assume Hvy: v < y.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is mul_SNo_Lt x y u v Hx Hy Hu Hv Hux Hvy.
Assume Hvy: v = y.
rewrite the current goal using Hvy (from left to right).
We will prove u * y + x * y x * y + u * y.
rewrite the current goal using add_SNo_com (u * y) (x * y) (SNo_mul_SNo u y Hu Hy) (SNo_mul_SNo x y Hx Hy) (from left to right).
We will prove x * y + u * y x * y + u * y.
Apply SNoLe_ref to the current goal.
Assume Hux: u = x.
rewrite the current goal using Hux (from left to right).
We will prove x * y + x * v x * y + x * v.
Apply SNoLe_ref to the current goal.
Theorem. (mul_SNo_SNoL_interpolate) The following is provable:
∀x y, SNo xSNo yuSNoL (x * y), (vSNoL x, wSNoL y, u + v * w v * y + x * w) (vSNoR x, wSNoR y, u + v * w v * y + x * w)
Proof:
Let x and y be given.
Assume Hx Hy.
Set P1 to be the term λu ⇒ vSNoL x, wSNoL y, u + v * w v * y + x * w of type setprop.
Set P2 to be the term λu ⇒ vSNoR x, wSNoR y, u + v * w v * y + x * w of type setprop.
Set P to be the term λu ⇒ P1 u P2 u of type setprop.
We prove the intermediate claim Lxy: SNo (x * y).
An exact proof term for the current goal is SNo_mul_SNo x y Hx Hy.
We prove the intermediate claim LI: ∀u, SNo uSNoLev u SNoLev (x * y)u < x * yP u.
Apply SNoLev_ind to the current goal.
Let u be given.
Assume Hu1: SNo u.
Assume IH: zSNoS_ (SNoLev u), SNoLev z SNoLev (x * y)z < x * yP z.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: u < x * y.
Apply dneg to the current goal.
Assume HNC: ¬ P u.
We prove the intermediate claim L1: vSNoL x, wSNoL y, v * y + x * w < u + v * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply SNoLtLe_or (v * y + x * w) (u + v * w) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Hv1 Hw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
We prove the intermediate claim L2: vSNoR x, wSNoR y, v * y + x * w < u + v * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply SNoLtLe_or (v * y + x * w) (u + v * w) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Hv1 Hw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
Apply SNoLt_irref u to the current goal.
Apply SNoLtLe_tra u (x * y) u Hu1 Lxy Hu1 Hu3 to the current goal.
We will prove x * y u.
Apply mul_SNo_eq_3 x y Hx Hy to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2.
Assume HE: x * y = SNoCut L R.
rewrite the current goal using HE (from left to right).
We will prove SNoCut L R u.
rewrite the current goal using SNo_eta u Hu1 (from left to right).
We will prove SNoCut L R SNoCut (SNoL u) (SNoR u).
Apply SNoCut_Le L R (SNoL u) (SNoR u) HLR (SNoCutP_SNoL_SNoR u Hu1) to the current goal.
We will prove zL, z < SNoCut (SNoL u) (SNoR u).
Let z be given.
Assume Hz.
rewrite the current goal using SNo_eta u Hu1 (from right to left).
We will prove z < u.
Apply HLE z Hz to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoL y.
Assume Hze: z = v * y + x * w + - v * w.
rewrite the current goal using Hze (from left to right).
We will prove v * y + x * w + - v * w < u.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply add_SNo_minus_Lt1b3 (v * y) (x * w) (v * w) u (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v w Hv1 Hw1) Hu1 to the current goal.
We will prove v * y + x * w < u + v * w.
An exact proof term for the current goal is L1 v Hv w Hw.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoR y.
Assume Hze: z = v * y + x * w + - v * w.
rewrite the current goal using Hze (from left to right).
We will prove v * y + x * w + - v * w < u.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply add_SNo_minus_Lt1b3 (v * y) (x * w) (v * w) u (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v w Hv1 Hw1) Hu1 to the current goal.
We will prove v * y + x * w < u + v * w.
An exact proof term for the current goal is L2 v Hv w Hw.
We will prove zSNoR u, SNoCut L R < z.
Let z be given.
Assume Hz: z SNoR u.
rewrite the current goal using HE (from right to left).
We will prove x * y < z.
Apply SNoR_E u Hu1 z Hz to the current goal.
Assume Hz1: SNo z.
Assume Hz2: SNoLev z SNoLev u.
Assume Hz3: u < z.
Apply SNoLt_trichotomy_or_impred z (x * y) Hz1 Lxy to the current goal.
Assume H1: z < x * y.
We prove the intermediate claim LPz: P z.
Apply IH z to the current goal.
We will prove z SNoS_ (SNoLev u).
Apply SNoS_I2 to the current goal.
An exact proof term for the current goal is Hz1.
An exact proof term for the current goal is Hu1.
An exact proof term for the current goal is Hz2.
We will prove SNoLev z SNoLev (x * y).
An exact proof term for the current goal is ordinal_TransSet (SNoLev (x * y)) (SNoLev_ordinal (x * y) Lxy) (SNoLev u) Hu2 (SNoLev z) Hz2.
We will prove z < x * y.
An exact proof term for the current goal is H1.
Apply LPz to the current goal.
Assume H2: P1 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v SNoL x.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w SNoL y.
Assume Hvw: z + v * w v * y + x * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim L3: z + v * w < u + v * w.
Apply SNoLeLt_tra (z + v * w) (v * y + x * w) (u + v * w) (SNo_add_SNo z (v * w) Hz1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) (SNo_add_SNo u (v * w) Hu1 Lvw) Hvw to the current goal.
We will prove v * y + x * w < u + v * w.
An exact proof term for the current goal is L1 v Hv w Hw.
We prove the intermediate claim L4: z < u.
An exact proof term for the current goal is add_SNo_Lt1_cancel z (v * w) u Hz1 Lvw Hu1 L3.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L4.
Assume H2: P2 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v SNoR x.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w SNoR y.
Assume Hvw: z + v * w v * y + x * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim L5: z + v * w < u + v * w.
Apply SNoLeLt_tra (z + v * w) (v * y + x * w) (u + v * w) (SNo_add_SNo z (v * w) Hz1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) (SNo_add_SNo u (v * w) Hu1 Lvw) Hvw to the current goal.
We will prove v * y + x * w < u + v * w.
An exact proof term for the current goal is L2 v Hv w Hw.
We prove the intermediate claim L6: z < u.
An exact proof term for the current goal is add_SNo_Lt1_cancel z (v * w) u Hz1 Lvw Hu1 L5.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L6.
Assume H1: z = x * y.
Apply In_no2cycle (SNoLev u) (SNoLev (x * y)) Hu2 to the current goal.
We will prove SNoLev (x * y) SNoLev u.
rewrite the current goal using H1 (from right to left).
An exact proof term for the current goal is Hz2.
Assume H1: x * y < z.
An exact proof term for the current goal is H1.
Let u be given.
Assume Hu: u SNoL (x * y).
Apply SNoL_E (x * y) Lxy u Hu to the current goal.
Assume Hu1: SNo u.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: u < x * y.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3.
Theorem. (mul_SNo_SNoL_interpolate_impred) The following is provable:
∀x y, SNo xSNo yuSNoL (x * y), ∀p : prop, (vSNoL x, wSNoL y, u + v * w v * y + x * wp)(vSNoR x, wSNoR y, u + v * w v * y + x * wp)p
Proof:
Let x and y be given.
Assume Hx Hy.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp1 Hp2.
Apply mul_SNo_SNoL_interpolate x y Hx Hy u Hu to the current goal.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw.
Theorem. (mul_SNo_SNoR_interpolate) The following is provable:
∀x y, SNo xSNo yuSNoR (x * y), (vSNoL x, wSNoR y, v * y + x * w u + v * w) (vSNoR x, wSNoL y, v * y + x * w u + v * w)
Proof:
Let x and y be given.
Assume Hx Hy.
Set P1 to be the term λu ⇒ vSNoL x, wSNoR y, v * y + x * w u + v * w of type setprop.
Set P2 to be the term λu ⇒ vSNoR x, wSNoL y, v * y + x * w u + v * w of type setprop.
Set P to be the term λu ⇒ P1 u P2 u of type setprop.
We prove the intermediate claim Lxy: SNo (x * y).
An exact proof term for the current goal is SNo_mul_SNo x y Hx Hy.
We prove the intermediate claim LI: ∀u, SNo uSNoLev u SNoLev (x * y)x * y < uP u.
Apply SNoLev_ind to the current goal.
Let u be given.
Assume Hu1: SNo u.
Assume IH: zSNoS_ (SNoLev u), SNoLev z SNoLev (x * y)x * y < zP z.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: x * y < u.
Apply dneg to the current goal.
Assume HNC: ¬ P u.
We prove the intermediate claim L1: vSNoL x, wSNoR y, u + v * w < v * y + x * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply SNoLtLe_or (u + v * w) (v * y + x * w) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Hv1 Hw1)) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
We prove the intermediate claim L2: vSNoR x, wSNoL y, u + v * w < v * y + x * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply SNoLtLe_or (u + v * w) (v * y + x * w) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Hv1 Hw1)) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
Apply SNoLt_irref (x * y) to the current goal.
Apply SNoLtLe_tra (x * y) u (x * y) Lxy Hu1 Lxy Hu3 to the current goal.
We will prove u x * y.
Apply mul_SNo_eq_3 x y Hx Hy to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2.
Assume HE: x * y = SNoCut L R.
rewrite the current goal using HE (from left to right).
We will prove u SNoCut L R.
rewrite the current goal using SNo_eta u Hu1 (from left to right).
We will prove SNoCut (SNoL u) (SNoR u) SNoCut L R.
Apply SNoCut_Le (SNoL u) (SNoR u) L R (SNoCutP_SNoL_SNoR u Hu1) HLR to the current goal.
We will prove zSNoL u, z < SNoCut L R.
Let z be given.
Assume Hz: z SNoL u.
rewrite the current goal using HE (from right to left).
We will prove z < x * y.
Apply SNoL_E u Hu1 z Hz to the current goal.
Assume Hz1: SNo z.
Assume Hz2: SNoLev z SNoLev u.
Assume Hz3: z < u.
Apply SNoLt_trichotomy_or_impred z (x * y) Hz1 Lxy to the current goal.
Assume H1: z < x * y.
An exact proof term for the current goal is H1.
Assume H1: z = x * y.
Apply In_no2cycle (SNoLev u) (SNoLev (x * y)) Hu2 to the current goal.
We will prove SNoLev (x * y) SNoLev u.
rewrite the current goal using H1 (from right to left).
An exact proof term for the current goal is Hz2.
Assume H1: x * y < z.
We prove the intermediate claim LPz: P z.
Apply IH z to the current goal.
We will prove z SNoS_ (SNoLev u).
Apply SNoS_I2 to the current goal.
An exact proof term for the current goal is Hz1.
An exact proof term for the current goal is Hu1.
An exact proof term for the current goal is Hz2.
We will prove SNoLev z SNoLev (x * y).
An exact proof term for the current goal is ordinal_TransSet (SNoLev (x * y)) (SNoLev_ordinal (x * y) Lxy) (SNoLev u) Hu2 (SNoLev z) Hz2.
We will prove x * y < z.
An exact proof term for the current goal is H1.
Apply LPz to the current goal.
Assume H2: P1 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v SNoL x.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w SNoR y.
Assume Hvw: v * y + x * w z + v * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim L3: u + v * w < z + v * w.
Apply SNoLtLe_tra (u + v * w) (v * y + x * w) (z + v * w) (SNo_add_SNo u (v * w) Hu1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) (SNo_add_SNo z (v * w) Hz1 Lvw) to the current goal.
We will prove u + v * w < v * y + x * w.
An exact proof term for the current goal is L1 v Hv w Hw.
An exact proof term for the current goal is Hvw.
We prove the intermediate claim L4: u < z.
An exact proof term for the current goal is add_SNo_Lt1_cancel u (v * w) z Hu1 Lvw Hz1 L3.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 L4 Hz3.
Assume H2: P2 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v SNoR x.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w SNoL y.
Assume Hvw: v * y + x * w z + v * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim L5: u + v * w < z + v * w.
Apply SNoLtLe_tra (u + v * w) (v * y + x * w) (z + v * w) (SNo_add_SNo u (v * w) Hu1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1)) (SNo_add_SNo z (v * w) Hz1 Lvw) to the current goal.
We will prove u + v * w < v * y + x * w.
An exact proof term for the current goal is L2 v Hv w Hw.
An exact proof term for the current goal is Hvw.
We prove the intermediate claim L6: u < z.
An exact proof term for the current goal is add_SNo_Lt1_cancel u (v * w) z Hu1 Lvw Hz1 L5.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 L6 Hz3.
We will prove zR, SNoCut (SNoL u) (SNoR u) < z.
Let z be given.
Assume Hz.
rewrite the current goal using SNo_eta u Hu1 (from right to left).
We will prove u < z.
Apply HRE z Hz to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoR y.
Assume Hze: z = v * y + x * w + - v * w.
rewrite the current goal using Hze (from left to right).
We will prove u < v * y + x * w + - v * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply add_SNo_minus_Lt2b3 (v * y) (x * w) (v * w) u (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v w Hv1 Hw1) Hu1 to the current goal.
We will prove u + v * w < v * y + x * w.
An exact proof term for the current goal is L1 v Hv w Hw.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoL y.
Assume Hze: z = v * y + x * w + - v * w.
rewrite the current goal using Hze (from left to right).
We will prove u < v * y + x * w + - v * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply add_SNo_minus_Lt2b3 (v * y) (x * w) (v * w) u (SNo_mul_SNo v y Hv1 Hy) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v w Hv1 Hw1) Hu1 to the current goal.
We will prove u + v * w < v * y + x * w.
An exact proof term for the current goal is L2 v Hv w Hw.
Let u be given.
Assume Hu: u SNoR (x * y).
Apply SNoR_E (x * y) Lxy u Hu to the current goal.
Assume Hu1: SNo u.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: x * y < u.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3.
Theorem. (mul_SNo_SNoR_interpolate_impred) The following is provable:
∀x y, SNo xSNo yuSNoR (x * y), ∀p : prop, (vSNoL x, wSNoR y, v * y + x * w u + v * wp)(vSNoR x, wSNoL y, v * y + x * w u + v * wp)p
Proof:
Let x and y be given.
Assume Hx Hy.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp1 Hp2.
Apply mul_SNo_SNoR_interpolate x y Hx Hy u Hu to the current goal.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw.
Theorem. (mul_SNo_Subq_lem) The following is provable:
∀x y X Y Z W, ∀U U', (∀u, u U(∀q : prop, (w0X, w1Y, u = w0 * y + x * w1 + - w0 * w1q)(z0Z, z1W, u = z0 * y + x * z1 + - z0 * z1q)q))(w0X, w1Y, w0 * y + x * w1 + - w0 * w1 U')(w0Z, w1W, w0 * y + x * w1 + - w0 * w1 U')U U'
Proof:
Let x, y, X, Y, Z, W, U and U' be given.
Assume HUE HU'I1 HU'I2.
Let u be given.
Assume Hu: u U.
Apply HUE u Hu to the current goal.
Let w be given.
Assume Hw.
Let z be given.
Assume Hz Huwz.
rewrite the current goal using Huwz (from left to right).
Apply HU'I1 to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is Hz.
Let w be given.
Assume Hw.
Let z be given.
Assume Hz Huwz.
rewrite the current goal using Huwz (from left to right).
Apply HU'I2 to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is Hz.
Theorem. (mul_SNo_zeroR) The following is provable:
∀x, SNo xx * 0 = 0
Proof:
Let x be given.
Assume Hx: SNo x.
Apply mul_SNo_eq_2 x 0 Hx SNo_0 to the current goal.
Let L and R be given.
Assume HLE HLI1 HLI2 HRE HRI1 HRI2 Hx0.
We will prove x * 0 = 0.
rewrite the current goal using Hx0 (from left to right).
We will prove SNoCut L R = 0.
We prove the intermediate claim LL0: L = 0.
Apply Empty_Subq_eq to the current goal.
Let w be given.
Assume Hw: w L.
We will prove False.
Apply HLE w Hw to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
rewrite the current goal using SNoL_0 (from left to right).
Assume Hv: v 0.
We will prove False.
An exact proof term for the current goal is EmptyE v Hv.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
rewrite the current goal using SNoR_0 (from left to right).
Assume Hv: v 0.
We will prove False.
An exact proof term for the current goal is EmptyE v Hv.
We prove the intermediate claim LR0: R = 0.
Apply Empty_Subq_eq to the current goal.
Let w be given.
Assume Hw: w R.
We will prove False.
Apply HRE w Hw to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
rewrite the current goal using SNoR_0 (from left to right).
Assume Hv: v 0.
We will prove False.
An exact proof term for the current goal is EmptyE v Hv.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
rewrite the current goal using SNoL_0 (from left to right).
Assume Hv: v 0.
We will prove False.
An exact proof term for the current goal is EmptyE v Hv.
rewrite the current goal using LL0 (from left to right).
rewrite the current goal using LR0 (from left to right).
An exact proof term for the current goal is SNoCut_0_0.
Theorem. (mul_SNo_oneR) The following is provable:
∀x, SNo xx * 1 = x
Proof:
Apply SNoLev_ind to the current goal.
Let x be given.
Assume Hx.
Assume IHx: wSNoS_ (SNoLev x), w * 1 = w.
Apply mul_SNo_eq_3 x 1 Hx SNo_1 to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 Hx1e.
We will prove x * 1 = x.
Apply mul_SNo_prop_1 x Hx 1 SNo_1 to the current goal.
Assume Hx1: SNo (x * 1).
Assume Hx11: uSNoL x, vSNoL 1, u * 1 + x * v < x * 1 + u * v.
Assume Hx12: uSNoR x, vSNoR 1, u * 1 + x * v < x * 1 + u * v.
Assume Hx13: uSNoL x, vSNoR 1, x * 1 + u * v < u * 1 + x * v.
Assume Hx14: uSNoR x, vSNoL 1, x * 1 + u * v < u * 1 + x * v.
We prove the intermediate claim L0L1: 0 SNoL 1.
rewrite the current goal using SNoL_1 (from left to right).
An exact proof term for the current goal is In_0_1.
rewrite the current goal using Hx1e (from left to right).
We will prove SNoCut L R = x.
rewrite the current goal using SNo_eta x Hx (from left to right).
We will prove SNoCut L R = SNoCut (SNoL x) (SNoR x).
Apply SNoCut_ext L R (SNoL x) (SNoR x) HLR (SNoCutP_SNoL_SNoR x Hx) to the current goal.
We will prove wL, w < SNoCut (SNoL x) (SNoR x).
Let w be given.
Assume Hw.
Apply SNoCutP_SNoCut_L (SNoL x) (SNoR x) (SNoCutP_SNoL_SNoR x Hx) to the current goal.
We will prove w SNoL x.
Apply HLE w Hw to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
rewrite the current goal using SNoL_1 (from left to right).
Assume Hv: v 1.
Apply cases_1 v Hv to the current goal.
Assume Hwuv: w = u * 1 + x * 0 + - u * 0.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _.
We prove the intermediate claim L1: w = u.
rewrite the current goal using Hwuv (from left to right).
Use transitivity with u * 1 + 0, and u * 1.
We will prove u * 1 + x * 0 + - u * 0 = u * 1 + 0.
Use f_equal.
We will prove x * 0 + - u * 0 = 0.
Use transitivity with x * 0 + 0, and x * 0.
We will prove x * 0 + - u * 0 = x * 0 + 0.
Use f_equal.
We will prove - u * 0 = 0.
Use transitivity with and - 0.
We will prove - u * 0 = - 0.
Use f_equal.
An exact proof term for the current goal is mul_SNo_zeroR u Hua.
An exact proof term for the current goal is minus_SNo_0.
We will prove x * 0 + 0 = x * 0.
An exact proof term for the current goal is add_SNo_0R (x * 0) (SNo_mul_SNo x 0 Hx SNo_0).
We will prove x * 0 = 0.
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
We will prove u * 1 + 0 = u * 1.
An exact proof term for the current goal is add_SNo_0R (u * 1) (SNo_mul_SNo u 1 Hua SNo_1).
We will prove u * 1 = u.
An exact proof term for the current goal is IHx u (SNoL_SNoS x Hx u Hu).
rewrite the current goal using L1 (from left to right).
An exact proof term for the current goal is Hu.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
rewrite the current goal using SNoR_1 (from left to right).
Assume Hv: v 0.
We will prove False.
An exact proof term for the current goal is EmptyE v Hv.
We will prove zR, SNoCut (SNoL x) (SNoR x) < z.
Let z be given.
Assume Hz.
Apply SNoCutP_SNoCut_R (SNoL x) (SNoR x) (SNoCutP_SNoL_SNoR x Hx) to the current goal.
We will prove z SNoR x.
Apply HRE z Hz to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
rewrite the current goal using SNoR_1 (from left to right).
Assume Hv: v 0.
We will prove False.
An exact proof term for the current goal is EmptyE v Hv.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
rewrite the current goal using SNoL_1 (from left to right).
Assume Hv: v 1.
Apply cases_1 v Hv to the current goal.
Assume Hzuv: z = u * 1 + x * 0 + - u * 0.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _.
We prove the intermediate claim L1: z = u.
rewrite the current goal using Hzuv (from left to right).
Use transitivity with u * 1 + 0, and u * 1.
We will prove u * 1 + x * 0 + - u * 0 = u * 1 + 0.
Use f_equal.
We will prove x * 0 + - u * 0 = 0.
Use transitivity with x * 0 + 0, and x * 0.
We will prove x * 0 + - u * 0 = x * 0 + 0.
Use f_equal.
We will prove - u * 0 = 0.
Use transitivity with and - 0.
We will prove - u * 0 = - 0.
Use f_equal.
An exact proof term for the current goal is mul_SNo_zeroR u Hua.
An exact proof term for the current goal is minus_SNo_0.
We will prove x * 0 + 0 = x * 0.
An exact proof term for the current goal is add_SNo_0R (x * 0) (SNo_mul_SNo x 0 Hx SNo_0).
We will prove x * 0 = 0.
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
We will prove u * 1 + 0 = u * 1.
An exact proof term for the current goal is add_SNo_0R (u * 1) (SNo_mul_SNo u 1 Hua SNo_1).
We will prove u * 1 = u.
An exact proof term for the current goal is IHx u (SNoR_SNoS x Hx u Hu).
rewrite the current goal using L1 (from left to right).
An exact proof term for the current goal is Hu.
We will prove wSNoL x, w < SNoCut L R.
Let w be given.
Assume Hw.
Apply SNoL_E x Hx w Hw to the current goal.
Assume Hwa _ _.
rewrite the current goal using Hx1e (from right to left).
We will prove w < x * 1.
We prove the intermediate claim L1: w * 1 + x * 0 = w.
Use transitivity with w * 1 + 0, and w * 1.
Use f_equal.
We will prove x * 0 = 0.
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
An exact proof term for the current goal is add_SNo_0R (w * 1) (SNo_mul_SNo w 1 Hwa SNo_1).
An exact proof term for the current goal is IHx w (SNoL_SNoS x Hx w Hw).
We prove the intermediate claim L2: x * 1 + w * 0 = x * 1.
Use transitivity with and x * 1 + 0.
Use f_equal.
An exact proof term for the current goal is mul_SNo_zeroR w Hwa.
An exact proof term for the current goal is add_SNo_0R (x * 1) (SNo_mul_SNo x 1 Hx SNo_1).
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
We will prove w * 1 + x * 0 < x * 1 + w * 0.
An exact proof term for the current goal is Hx11 w Hw 0 L0L1.
We will prove zSNoR x, SNoCut L R < z.
Let z be given.
Assume Hz.
Apply SNoR_E x Hx z Hz to the current goal.
Assume Hza _ _.
rewrite the current goal using Hx1e (from right to left).
We will prove x * 1 < z.
We prove the intermediate claim L1: x * 1 + z * 0 = x * 1.
Use transitivity with and x * 1 + 0.
Use f_equal.
An exact proof term for the current goal is mul_SNo_zeroR z Hza.
An exact proof term for the current goal is add_SNo_0R (x * 1) (SNo_mul_SNo x 1 Hx SNo_1).
We prove the intermediate claim L2: z * 1 + x * 0 = z.
Use transitivity with z * 1 + 0, and z * 1.
Use f_equal.
We will prove x * 0 = 0.
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
An exact proof term for the current goal is add_SNo_0R (z * 1) (SNo_mul_SNo z 1 Hza SNo_1).
An exact proof term for the current goal is IHx z (SNoR_SNoS x Hx z Hz).
rewrite the current goal using L2 (from right to left).
rewrite the current goal using L1 (from right to left).
An exact proof term for the current goal is Hx14 z Hz 0 L0L1.
Theorem. (mul_SNo_com) The following is provable:
∀x y, SNo xSNo yx * y = y * x
Proof:
Apply SNoLev_ind2 to the current goal.
Let x and y be given.
Assume Hx Hy.
Assume IHx: wSNoS_ (SNoLev x), w * y = y * w.
Assume IHy: zSNoS_ (SNoLev y), x * z = z * x.
Assume IHxy: wSNoS_ (SNoLev x), zSNoS_ (SNoLev y), w * z = z * w.
Apply mul_SNo_eq_3 x y Hx Hy to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 Hxye.
Apply mul_SNo_eq_3 y x Hy Hx to the current goal.
Let L' and R' be given.
Assume HL'R' HL'E HL'I1 HL'I2 HR'E HR'I1 HR'I2 Hyxe.
rewrite the current goal using Hxye (from left to right).
rewrite the current goal using Hyxe (from left to right).
We will prove SNoCut L R = SNoCut L' R'.
We prove the intermediate claim LLL': L = L'.
Apply set_ext to the current goal.
We will prove L L'.
Apply mul_SNo_Subq_lem x y (SNoL x) (SNoL y) (SNoR x) (SNoR y) L L' HLE to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
Assume Hv: v SNoL y.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoL_SNoS x Hx u Hu) (from left to right).
rewrite the current goal using IHy v (SNoL_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using IHxy u (SNoL_SNoS x Hx u Hu) v (SNoL_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * u) (v * x) (- v * u) (SNo_mul_SNo y u Hy Hua) (SNo_mul_SNo v x Hva Hx) (SNo_minus_SNo (v * u) (SNo_mul_SNo v u Hva Hua)) (from left to right).
Apply HL'I1 to the current goal.
An exact proof term for the current goal is Hv.
An exact proof term for the current goal is Hu.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
Assume Hv: v SNoR y.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoR_SNoS x Hx u Hu) (from left to right).
rewrite the current goal using IHy v (SNoR_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using IHxy u (SNoR_SNoS x Hx u Hu) v (SNoR_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * u) (v * x) (- v * u) (SNo_mul_SNo y u Hy Hua) (SNo_mul_SNo v x Hva Hx) (SNo_minus_SNo (v * u) (SNo_mul_SNo v u Hva Hua)) (from left to right).
Apply HL'I2 to the current goal.
An exact proof term for the current goal is Hv.
An exact proof term for the current goal is Hu.
We will prove L' L.
Apply mul_SNo_Subq_lem y x (SNoL y) (SNoL x) (SNoR y) (SNoR x) L' L HL'E to the current goal.
Let v be given.
Assume Hv: v SNoL y.
Let u be given.
Assume Hu: u SNoL x.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoL_SNoS x Hx u Hu) (from right to left).
rewrite the current goal using IHy v (SNoL_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using IHxy u (SNoL_SNoS x Hx u Hu) v (SNoL_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using add_SNo_com_3_0_1 (x * v) (u * y) (- u * v) (SNo_mul_SNo x v Hx Hva) (SNo_mul_SNo u y Hua Hy) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Apply HLI1 to the current goal.
An exact proof term for the current goal is Hu.
An exact proof term for the current goal is Hv.
Let v be given.
Assume Hv: v SNoR y.
Let u be given.
Assume Hu: u SNoR x.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoR_SNoS x Hx u Hu) (from right to left).
rewrite the current goal using IHy v (SNoR_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using IHxy u (SNoR_SNoS x Hx u Hu) v (SNoR_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using add_SNo_com_3_0_1 (x * v) (u * y) (- u * v) (SNo_mul_SNo x v Hx Hva) (SNo_mul_SNo u y Hua Hy) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Apply HLI2 to the current goal.
An exact proof term for the current goal is Hu.
An exact proof term for the current goal is Hv.
We prove the intermediate claim LRR': R = R'.
Apply set_ext to the current goal.
Apply mul_SNo_Subq_lem x y (SNoL x) (SNoR y) (SNoR x) (SNoL y) R R' HRE to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
Assume Hv: v SNoR y.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoL_SNoS x Hx u Hu) (from left to right).
rewrite the current goal using IHy v (SNoR_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using IHxy u (SNoL_SNoS x Hx u Hu) v (SNoR_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * u) (v * x) (- v * u) (SNo_mul_SNo y u Hy Hua) (SNo_mul_SNo v x Hva Hx) (SNo_minus_SNo (v * u) (SNo_mul_SNo v u Hva Hua)) (from left to right).
Apply HR'I2 to the current goal.
An exact proof term for the current goal is Hv.
An exact proof term for the current goal is Hu.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
Assume Hv: v SNoL y.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoR_SNoS x Hx u Hu) (from left to right).
rewrite the current goal using IHy v (SNoL_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using IHxy u (SNoR_SNoS x Hx u Hu) v (SNoL_SNoS y Hy v Hv) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * u) (v * x) (- v * u) (SNo_mul_SNo y u Hy Hua) (SNo_mul_SNo v x Hva Hx) (SNo_minus_SNo (v * u) (SNo_mul_SNo v u Hva Hua)) (from left to right).
Apply HR'I1 to the current goal.
An exact proof term for the current goal is Hv.
An exact proof term for the current goal is Hu.
Apply mul_SNo_Subq_lem y x (SNoL y) (SNoR x) (SNoR y) (SNoL x) R' R HR'E to the current goal.
Let v be given.
Assume Hv: v SNoL y.
Let u be given.
Assume Hu: u SNoR x.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoR_SNoS x Hx u Hu) (from right to left).
rewrite the current goal using IHy v (SNoL_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using IHxy u (SNoR_SNoS x Hx u Hu) v (SNoL_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using add_SNo_com_3_0_1 (x * v) (u * y) (- u * v) (SNo_mul_SNo x v Hx Hva) (SNo_mul_SNo u y Hua Hy) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Apply HRI2 to the current goal.
An exact proof term for the current goal is Hu.
An exact proof term for the current goal is Hv.
Let v be given.
Assume Hv: v SNoR y.
Let u be given.
Assume Hu: u SNoL x.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
rewrite the current goal using IHx u (SNoL_SNoS x Hx u Hu) (from right to left).
rewrite the current goal using IHy v (SNoR_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using IHxy u (SNoL_SNoS x Hx u Hu) v (SNoR_SNoS y Hy v Hv) (from right to left).
rewrite the current goal using add_SNo_com_3_0_1 (x * v) (u * y) (- u * v) (SNo_mul_SNo x v Hx Hva) (SNo_mul_SNo u y Hua Hy) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Apply HRI1 to the current goal.
An exact proof term for the current goal is Hu.
An exact proof term for the current goal is Hv.
rewrite the current goal using LLL' (from left to right).
rewrite the current goal using LRR' (from left to right).
Use reflexivity.
Theorem. (mul_SNo_minus_distrL) The following is provable:
∀x y, SNo xSNo y(- x) * y = - x * y
Proof:
Apply SNoLev_ind2 to the current goal.
Let x and y be given.
Assume Hx Hy.
Assume IHx: wSNoS_ (SNoLev x), (- w) * y = - w * y.
Assume IHy: zSNoS_ (SNoLev y), (- x) * z = - x * z.
Assume IHxy: wSNoS_ (SNoLev x), zSNoS_ (SNoLev y), (- w) * z = - w * z.
We prove the intermediate claim Lmx: SNo (- x).
An exact proof term for the current goal is SNo_minus_SNo x Hx.
Apply mul_SNo_eq_3 x y Hx Hy to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2.
Assume Hxye: x * y = SNoCut L R.
Apply mul_SNo_eq_3 (- x) y (SNo_minus_SNo x Hx) Hy to the current goal.
Let L' and R' be given.
Assume HL'R' HL'E HL'I1 HL'I2 HR'E HR'I1 HR'I2.
Assume Hmxye: (- x) * y = SNoCut L' R'.
We prove the intermediate claim L1: - (x * y) = SNoCut {- z|zR} {- w|wL}.
rewrite the current goal using Hxye (from left to right).
An exact proof term for the current goal is minus_SNoCut_eq L R HLR.
rewrite the current goal using L1 (from left to right).
rewrite the current goal using Hmxye (from left to right).
We will prove SNoCut L' R' = SNoCut {- z|zR} {- w|wL}.
Use f_equal.
We will prove L' = {- z|zR}.
Apply set_ext to the current goal.
Apply mul_SNo_Subq_lem (- x) y (SNoL (- x)) (SNoL y) (SNoR (- x)) (SNoR y) L' {- z|zR} HL'E to the current goal.
Let u be given.
Assume Hu: u SNoL (- x).
Let v be given.
Assume Hv: v SNoL y.
Apply SNoL_E (- x) Lmx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev (- x).
Assume Huc: u < - x.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoR x.
Apply SNoR_I x Hx (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev x.
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub.
We will prove x < - u.
An exact proof term for the current goal is minus_SNo_Lt_contra2 u x Hua Hx Huc.
We will prove u * y + (- x) * v + - u * v {- z|zR}.
We prove the intermediate claim L1: u * y + (- x) * v + - u * v = - ((- u) * y + x * v + - (- u) * v).
Use symmetry.
Use transitivity with and - (- u) * y + - x * v + - - (- u) * v.
An exact proof term for the current goal is minus_add_SNo_distr_3 ((- u) * y) (x * v) (- (- u) * v) (SNo_mul_SNo (- u) y Lmu1 Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo ((- u) * v) (SNo_mul_SNo (- u) v Lmu1 Hva)).
Use f_equal.
We will prove - (- u) * y = u * y.
Use transitivity with and (- - u) * y.
Use symmetry.
An exact proof term for the current goal is IHx (- u) (SNoR_SNoS x Hx (- u) Lmu2).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
Use f_equal.
We will prove - x * v = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoL_SNoS y Hy v Hv).
Use f_equal.
We will prove - (- u) * v = u * v.
Use transitivity with and (- - u) * v.
Use symmetry.
We will prove (- - u) * v = - (- u) * v.
An exact proof term for the current goal is IHxy (- u) (SNoR_SNoS x Hx (- u) Lmu2) v (SNoL_SNoS y Hy v Hv).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
rewrite the current goal using L1 (from left to right).
We will prove - ((- u) * y + x * v + - (- u) * v) {- z|zR}.
Apply ReplI R (λz ⇒ - z) to the current goal.
We will prove (- u) * y + x * v + - (- u) * v R.
Apply HRI2 to the current goal.
We will prove - u SNoR x.
An exact proof term for the current goal is Lmu2.
We will prove v SNoL y.
An exact proof term for the current goal is Hv.
Let u be given.
Assume Hu: u SNoR (- x).
Let v be given.
Assume Hv: v SNoR y.
Apply SNoR_E (- x) Lmx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev (- x).
Assume Huc: - x < u.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoL x.
Apply SNoL_I x Hx (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev x.
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub.
We will prove - u < x.
An exact proof term for the current goal is minus_SNo_Lt_contra1 x u Hx Hua Huc.
We will prove u * y + (- x) * v + - u * v {- z|zR}.
We prove the intermediate claim L1: u * y + (- x) * v + - u * v = - ((- u) * y + x * v + - (- u) * v).
Use symmetry.
Use transitivity with and - (- u) * y + - x * v + - - (- u) * v.
An exact proof term for the current goal is minus_add_SNo_distr_3 ((- u) * y) (x * v) (- (- u) * v) (SNo_mul_SNo (- u) y Lmu1 Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo ((- u) * v) (SNo_mul_SNo (- u) v Lmu1 Hva)).
Use f_equal.
We will prove - (- u) * y = u * y.
Use transitivity with and (- - u) * y.
Use symmetry.
An exact proof term for the current goal is IHx (- u) (SNoL_SNoS x Hx (- u) Lmu2).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
Use f_equal.
We will prove - x * v = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoR_SNoS y Hy v Hv).
Use f_equal.
We will prove - (- u) * v = u * v.
Use transitivity with and (- - u) * v.
Use symmetry.
We will prove (- - u) * v = - (- u) * v.
An exact proof term for the current goal is IHxy (- u) (SNoL_SNoS x Hx (- u) Lmu2) v (SNoR_SNoS y Hy v Hv).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
rewrite the current goal using L1 (from left to right).
We will prove - ((- u) * y + x * v + - (- u) * v) {- z|zR}.
Apply ReplI R (λz ⇒ - z) to the current goal.
We will prove (- u) * y + x * v + - (- u) * v R.
Apply HRI1 to the current goal.
We will prove - u SNoL x.
An exact proof term for the current goal is Lmu2.
We will prove v SNoR y.
An exact proof term for the current goal is Hv.
We will prove {- z|zR} L'.
Let a be given.
Assume Ha.
Apply ReplE_impred R (λz ⇒ - z) a Ha to the current goal.
Let z be given.
Assume Hz: z R.
Assume Hze: a = - z.
rewrite the current goal using Hze (from left to right).
We will prove - z L'.
Apply HRE z Hz to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
Assume Hv: v SNoR y.
Assume Hzuv: z = u * y + x * v + - u * v.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev x.
Assume Huc: u < x.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoR (- x).
Apply SNoR_I (- x) (SNo_minus_SNo x Hx) (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev (- x).
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub.
We will prove - x < - u.
An exact proof term for the current goal is minus_SNo_Lt_contra u x Hua Hx Huc.
We prove the intermediate claim L1: - z = (- u) * y + (- x) * v + - (- u) * v.
rewrite the current goal using Hzuv (from left to right).
rewrite the current goal using minus_add_SNo_distr_3 (u * y) (x * v) (- (u * v)) (SNo_mul_SNo u y Hua Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Use f_equal.
We will prove - (u * y) = (- u) * y.
Use symmetry.
An exact proof term for the current goal is IHx u (SNoL_SNoS x Hx u Hu).
Use f_equal.
We will prove - (x * v) = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoR_SNoS y Hy v Hv).
Use f_equal.
We will prove - (u * v) = (- u) * v.
Use symmetry.
An exact proof term for the current goal is IHxy u (SNoL_SNoS x Hx u Hu) v (SNoR_SNoS y Hy v Hv).
rewrite the current goal using L1 (from left to right).
We will prove (- u) * y + (- x) * v + - (- u) * v L'.
Apply HL'I2 to the current goal.
An exact proof term for the current goal is Lmu2.
An exact proof term for the current goal is Hv.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
Assume Hv: v SNoL y.
Assume Hzuv: z = u * y + x * v + - u * v.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev x.
Assume Huc: x < u.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoL (- x).
Apply SNoL_I (- x) (SNo_minus_SNo x Hx) (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev (- x).
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub.
We will prove - u < - x.
An exact proof term for the current goal is minus_SNo_Lt_contra x u Hx Hua Huc.
We prove the intermediate claim L1: - z = (- u) * y + (- x) * v + - (- u) * v.
rewrite the current goal using Hzuv (from left to right).
rewrite the current goal using minus_add_SNo_distr_3 (u * y) (x * v) (- (u * v)) (SNo_mul_SNo u y Hua Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Use f_equal.
We will prove - (u * y) = (- u) * y.
Use symmetry.
An exact proof term for the current goal is IHx u (SNoR_SNoS x Hx u Hu).
Use f_equal.
We will prove - (x * v) = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoL_SNoS y Hy v Hv).
Use f_equal.
We will prove - (u * v) = (- u) * v.
Use symmetry.
An exact proof term for the current goal is IHxy u (SNoR_SNoS x Hx u Hu) v (SNoL_SNoS y Hy v Hv).
rewrite the current goal using L1 (from left to right).
We will prove (- u) * y + (- x) * v + - (- u) * v L'.
Apply HL'I1 to the current goal.
An exact proof term for the current goal is Lmu2.
An exact proof term for the current goal is Hv.
We will prove R' = {- w|wL}.
Apply set_ext to the current goal.
Apply mul_SNo_Subq_lem (- x) y (SNoL (- x)) (SNoR y) (SNoR (- x)) (SNoL y) R' {- w|wL} HR'E to the current goal.
Let u be given.
Assume Hu: u SNoL (- x).
Let v be given.
Assume Hv: v SNoR y.
Apply SNoL_E (- x) Lmx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev (- x).
Assume Huc: u < - x.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoR x.
Apply SNoR_I x Hx (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev x.
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub.
We will prove x < - u.
An exact proof term for the current goal is minus_SNo_Lt_contra2 u x Hua Hx Huc.
We will prove u * y + (- x) * v + - u * v {- w|wL}.
We prove the intermediate claim L1: u * y + (- x) * v + - u * v = - ((- u) * y + x * v + - (- u) * v).
Use symmetry.
Use transitivity with and - (- u) * y + - x * v + - - (- u) * v.
An exact proof term for the current goal is minus_add_SNo_distr_3 ((- u) * y) (x * v) (- (- u) * v) (SNo_mul_SNo (- u) y Lmu1 Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo ((- u) * v) (SNo_mul_SNo (- u) v Lmu1 Hva)).
Use f_equal.
We will prove - (- u) * y = u * y.
Use transitivity with and (- - u) * y.
Use symmetry.
An exact proof term for the current goal is IHx (- u) (SNoR_SNoS x Hx (- u) Lmu2).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
Use f_equal.
We will prove - x * v = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoR_SNoS y Hy v Hv).
Use f_equal.
We will prove - (- u) * v = u * v.
Use transitivity with and (- - u) * v.
Use symmetry.
We will prove (- - u) * v = - (- u) * v.
An exact proof term for the current goal is IHxy (- u) (SNoR_SNoS x Hx (- u) Lmu2) v (SNoR_SNoS y Hy v Hv).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
rewrite the current goal using L1 (from left to right).
We will prove - ((- u) * y + x * v + - (- u) * v) {- w|wL}.
Apply ReplI L (λw ⇒ - w) to the current goal.
We will prove (- u) * y + x * v + - (- u) * v L.
Apply HLI2 to the current goal.
We will prove - u SNoR x.
An exact proof term for the current goal is Lmu2.
We will prove v SNoR y.
An exact proof term for the current goal is Hv.
Let u be given.
Assume Hu: u SNoR (- x).
Let v be given.
Assume Hv: v SNoL y.
Apply SNoR_E (- x) Lmx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev (- x).
Assume Huc: - x < u.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoL x.
Apply SNoL_I x Hx (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev x.
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub.
We will prove - u < x.
An exact proof term for the current goal is minus_SNo_Lt_contra1 x u Hx Hua Huc.
We will prove u * y + (- x) * v + - u * v {- w|wL}.
We prove the intermediate claim L1: u * y + (- x) * v + - u * v = - ((- u) * y + x * v + - (- u) * v).
Use symmetry.
Use transitivity with and - (- u) * y + - x * v + - - (- u) * v.
An exact proof term for the current goal is minus_add_SNo_distr_3 ((- u) * y) (x * v) (- (- u) * v) (SNo_mul_SNo (- u) y Lmu1 Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo ((- u) * v) (SNo_mul_SNo (- u) v Lmu1 Hva)).
Use f_equal.
We will prove - (- u) * y = u * y.
Use transitivity with and (- - u) * y.
Use symmetry.
An exact proof term for the current goal is IHx (- u) (SNoL_SNoS x Hx (- u) Lmu2).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
Use f_equal.
We will prove - x * v = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoL_SNoS y Hy v Hv).
Use f_equal.
We will prove - (- u) * v = u * v.
Use transitivity with and (- - u) * v.
Use symmetry.
We will prove (- - u) * v = - (- u) * v.
An exact proof term for the current goal is IHxy (- u) (SNoL_SNoS x Hx (- u) Lmu2) v (SNoL_SNoS y Hy v Hv).
rewrite the current goal using minus_SNo_invol u Hua (from left to right).
Use reflexivity.
rewrite the current goal using L1 (from left to right).
We will prove - ((- u) * y + x * v + - (- u) * v) {- w|wL}.
Apply ReplI L (λw ⇒ - w) to the current goal.
We will prove (- u) * y + x * v + - (- u) * v L.
Apply HLI1 to the current goal.
We will prove - u SNoL x.
An exact proof term for the current goal is Lmu2.
We will prove v SNoL y.
An exact proof term for the current goal is Hv.
We will prove {- w|wL} R'.
Let a be given.
Assume Ha.
Apply ReplE_impred L (λw ⇒ - w) a Ha to the current goal.
Let w be given.
Assume Hw: w L.
Assume Hwe: a = - w.
rewrite the current goal using Hwe (from left to right).
We will prove - w R'.
Apply HLE w Hw to the current goal.
Let u be given.
Assume Hu: u SNoL x.
Let v be given.
Assume Hv: v SNoL y.
Assume Hwuv: w = u * y + x * v + - u * v.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev x.
Assume Huc: u < x.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoR (- x).
Apply SNoR_I (- x) (SNo_minus_SNo x Hx) (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev (- x).
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub.
We will prove - x < - u.
An exact proof term for the current goal is minus_SNo_Lt_contra u x Hua Hx Huc.
We prove the intermediate claim L1: - w = (- u) * y + (- x) * v + - (- u) * v.
rewrite the current goal using Hwuv (from left to right).
rewrite the current goal using minus_add_SNo_distr_3 (u * y) (x * v) (- (u * v)) (SNo_mul_SNo u y Hua Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Use f_equal.
We will prove - (u * y) = (- u) * y.
Use symmetry.
An exact proof term for the current goal is IHx u (SNoL_SNoS x Hx u Hu).
Use f_equal.
We will prove - (x * v) = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoL_SNoS y Hy v Hv).
Use f_equal.
We will prove - (u * v) = (- u) * v.
Use symmetry.
An exact proof term for the current goal is IHxy u (SNoL_SNoS x Hx u Hu) v (SNoL_SNoS y Hy v Hv).
rewrite the current goal using L1 (from left to right).
We will prove (- u) * y + (- x) * v + - (- u) * v R'.
Apply HR'I2 to the current goal.
An exact proof term for the current goal is Lmu2.
An exact proof term for the current goal is Hv.
Let u be given.
Assume Hu: u SNoR x.
Let v be given.
Assume Hv: v SNoR y.
Assume Hwuv: w = u * y + x * v + - u * v.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua: SNo u.
Assume Hub: SNoLev u SNoLev x.
Assume Huc: x < u.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _.
We prove the intermediate claim Lmu1: SNo (- u).
An exact proof term for the current goal is SNo_minus_SNo u Hua.
We prove the intermediate claim Lmu2: - u SNoL (- x).
Apply SNoL_I (- x) (SNo_minus_SNo x Hx) (- u) Lmu1 to the current goal.
We will prove SNoLev (- u) SNoLev (- x).
rewrite the current goal using minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub.
We will prove - u < - x.
An exact proof term for the current goal is minus_SNo_Lt_contra x u Hx Hua Huc.
We prove the intermediate claim L1: - w = (- u) * y + (- x) * v + - (- u) * v.
rewrite the current goal using Hwuv (from left to right).
rewrite the current goal using minus_add_SNo_distr_3 (u * y) (x * v) (- (u * v)) (SNo_mul_SNo u y Hua Hy) (SNo_mul_SNo x v Hx Hva) (SNo_minus_SNo (u * v) (SNo_mul_SNo u v Hua Hva)) (from left to right).
Use f_equal.
We will prove - (u * y) = (- u) * y.
Use symmetry.
An exact proof term for the current goal is IHx u (SNoR_SNoS x Hx u Hu).
Use f_equal.
We will prove - (x * v) = (- x) * v.
Use symmetry.
An exact proof term for the current goal is IHy v (SNoR_SNoS y Hy v Hv).
Use f_equal.
We will prove - (u * v) = (- u) * v.
Use symmetry.
An exact proof term for the current goal is IHxy u (SNoR_SNoS x Hx u Hu) v (SNoR_SNoS y Hy v Hv).
rewrite the current goal using L1 (from left to right).
We will prove (- u) * y + (- x) * v + - (- u) * v R'.
Apply HR'I1 to the current goal.
An exact proof term for the current goal is Lmu2.
An exact proof term for the current goal is Hv.
Theorem. (mul_SNo_minus_distrR) The following is provable:
∀x y, SNo xSNo yx * (- y) = - (x * y)
Proof:
Let x and y be given.
Assume Hx Hy.
rewrite the current goal using mul_SNo_com x y Hx Hy (from left to right).
rewrite the current goal using mul_SNo_com x (- y) Hx (SNo_minus_SNo y Hy) (from left to right).
An exact proof term for the current goal is mul_SNo_minus_distrL y x Hy Hx.
Theorem. (mul_SNo_distrR) The following is provable:
∀x y z, SNo xSNo ySNo z(x + y) * z = x * z + y * z
Proof:
Set P to be the term λx y z ⇒ (x + y) * z = x * z + y * z of type setsetsetprop.
We will prove ∀x y z, SNo xSNo ySNo zP x y z.
Apply SNoLev_ind3 P to the current goal.
Let x, y and z be given.
Assume Hx Hy Hz.
Assume IHx: uSNoS_ (SNoLev x), (u + y) * z = u * z + y * z.
Assume IHy: vSNoS_ (SNoLev y), (x + v) * z = x * z + v * z.
Assume IHz: wSNoS_ (SNoLev z), (x + y) * w = x * w + y * w.
Assume IHxy: uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), (u + v) * z = u * z + v * z.
Assume IHxz: uSNoS_ (SNoLev x), wSNoS_ (SNoLev z), (u + y) * w = u * w + y * w.
Assume IHyz: vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), (x + v) * w = x * w + v * w.
Assume IHxyz: uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), (u + v) * w = u * w + v * w.
We will prove (x + y) * z = x * z + y * z.
Apply mul_SNo_eq_3 (x + y) z (SNo_add_SNo x y Hx Hy) Hz to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2.
Assume HE: (x + y) * z = SNoCut L R.
Set L1 to be the term {w + y * z|wSNoL (x * z)}.
Set L2 to be the term {x * z + w|wSNoL (y * z)}.
Set R1 to be the term {w + y * z|wSNoR (x * z)}.
Set R2 to be the term {x * z + w|wSNoR (y * z)}.
We prove the intermediate claim Lxy: SNo (x + y).
An exact proof term for the current goal is SNo_add_SNo x y Hx Hy.
We prove the intermediate claim Lxyz: SNo ((x + y) * z).
An exact proof term for the current goal is SNo_mul_SNo (x + y) z Lxy Hz.
We prove the intermediate claim Lxz: SNo (x * z).
An exact proof term for the current goal is SNo_mul_SNo x z Hx Hz.
We prove the intermediate claim Lyz: SNo (y * z).
An exact proof term for the current goal is SNo_mul_SNo y z Hy Hz.
We prove the intermediate claim Lxzyz: SNo (x * z + y * z).
An exact proof term for the current goal is SNo_add_SNo (x * z) (y * z) Lxz Lyz.
We prove the intermediate claim LE: x * z + y * z = SNoCut (L1 L2) (R1 R2).
An exact proof term for the current goal is add_SNo_eq (x * z) (SNo_mul_SNo x z Hx Hz) (y * z) (SNo_mul_SNo y z Hy Hz).
rewrite the current goal using HE (from left to right).
rewrite the current goal using LE (from left to right).
Apply SNoCut_ext to the current goal.
An exact proof term for the current goal is HLR.
An exact proof term for the current goal is add_SNo_SNoCutP (x * z) (y * z) (SNo_mul_SNo x z Hx Hz) (SNo_mul_SNo y z Hy Hz).
Let u be given.
Assume Hu: u L.
rewrite the current goal using LE (from right to left).
We will prove u < x * z + y * z.
Apply HLE u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL (x + y).
Let w be given.
Assume Hw: w SNoL z.
Assume Hue: u = v * z + (x + y) * w + - v * w.
rewrite the current goal using Hue (from left to right).
We will prove v * z + (x + y) * w + - v * w < x * z + y * z.
Apply SNoL_E (x + y) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lxwyw: SNo (x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (x * w) (y * w) Lxw Lyw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvzxwyw: SNo (v * z + x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo_3 (v * z) (x * w) (y * w) Lvz Lxw Lyw.
We prove the intermediate claim Lxzyzvw: SNo (x * z + y * z + v * w).
An exact proof term for the current goal is SNo_add_SNo_3 (x * z) (y * z) (v * w) Lxz Lyz Lvw.
Apply add_SNo_minus_Lt1b3 (v * z) ((x + y) * w) (v * w) (x * z + y * z) Lvz Lxyw Lvw Lxzyz to the current goal.
We will prove v * z + (x + y) * w < (x * z + y * z) + v * w.
rewrite the current goal using IHz w (SNoL_SNoS z Hz w Hw) (from left to right).
We will prove v * z + x * w + y * w < (x * z + y * z) + v * w.
rewrite the current goal using add_SNo_assoc (x * z) (y * z) (v * w) Lxz Lyz Lvw (from right to left).
We will prove v * z + x * w + y * w < x * z + y * z + v * w.
Apply add_SNo_SNoL_interpolate x y Hx Hy v Hv to the current goal.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoL x.
Assume Hvu: v u + y.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Luy: SNo (u + y).
An exact proof term for the current goal is SNo_add_SNo u y Hu1 Hy.
Apply add_SNo_Lt1_cancel (v * z + x * w + y * w) (u * w) (x * z + y * z + v * w) Lvzxwyw Luw Lxzyzvw to the current goal.
We will prove (v * z + x * w + y * w) + u * w < (x * z + y * z + v * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w.
Apply SNoLeLt_tra (v * z + x * w + y * w + u * w) (u * z + y * z + v * w + x * w) (x * z + y * z + v * w + u * w) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) (SNo_add_SNo_4 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) to the current goal.
We will prove v * z + x * w + y * w + u * w u * z + y * z + v * w + x * w.
rewrite the current goal using add_SNo_com_3_0_1 (v * z) (x * w) (y * w + u * w) Lvz Lxw (SNo_add_SNo (y * w) (u * w) Lyw Luw) (from left to right).
We will prove x * w + v * z + y * w + u * w u * z + y * z + v * w + x * w.
rewrite the current goal using add_SNo_rotate_4_1 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw (from left to right).
We will prove x * w + v * z + y * w + u * w x * w + u * z + y * z + v * w.
Apply add_SNo_Le2 (x * w) (v * z + y * w + u * w) (u * z + y * z + v * w) Lxw (SNo_add_SNo_3 (v * z) (y * w) (u * w) Lvz Lyw Luw) (SNo_add_SNo_3 (u * z) (y * z) (v * w) Luz Lyz Lvw) to the current goal.
We will prove v * z + y * w + u * w u * z + y * z + v * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove v * z + u * w + y * w u * z + y * z + v * w.
rewrite the current goal using IHxz u (SNoL_SNoS x Hx u Hu) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove v * z + (u + y) * w u * z + y * z + v * w.
rewrite the current goal using add_SNo_assoc (u * z) (y * z) (v * w) Luz Lyz Lvw (from left to right).
We will prove v * z + (u + y) * w (u * z + y * z) + v * w.
rewrite the current goal using IHx u (SNoL_SNoS x Hx u Hu) (from right to left).
We will prove v * z + (u + y) * w (u + y) * z + v * w.
Apply mul_SNo_Le (u + y) z v w Luy Hz Hv1 Hw1 to the current goal.
We will prove v u + y.
An exact proof term for the current goal is Hvu.
We will prove w z.
Apply SNoLtLe to the current goal.
We will prove w < z.
An exact proof term for the current goal is Hw3.
We will prove u * z + y * z + v * w + x * w < x * z + y * z + v * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (y * z) (v * w + x * w) Luz Lyz (SNo_add_SNo (v * w) (x * w) Lvw Lxw) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (y * z) (v * w + u * w) Lxz Lyz (SNo_add_SNo (v * w) (u * w) Lvw Luw) (from left to right).
We will prove y * z + u * z + v * w + x * w < y * z + x * z + v * w + u * w.
Apply add_SNo_Lt2 (y * z) (u * z + v * w + x * w) (x * z + v * w + u * w) Lyz (SNo_add_SNo_3 (u * z) (v * w) (x * w) Luz Lvw Lxw) (SNo_add_SNo_3 (x * z) (v * w) (u * w) Lxz Lvw Luw) to the current goal.
We will prove u * z + v * w + x * w < x * z + v * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (x * w) Luz Lvw Lxw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (v * w) (u * w) Lxz Lvw Luw (from left to right).
We will prove v * w + u * z + x * w < v * w + x * z + u * w.
Apply add_SNo_Lt2 (v * w) (u * z + x * w) (x * z + u * w) Lvw (SNo_add_SNo (u * z) (x * w) Luz Lxw) (SNo_add_SNo (x * z) (u * w) Lxz Luw) to the current goal.
We will prove u * z + x * w < x * z + u * w.
An exact proof term for the current goal is mul_SNo_Lt x z u w Hx Hz Hu1 Hw1 Hu3 Hw3.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoL y.
Assume Hvu: v x + u.
Apply SNoL_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Lxu: SNo (x + u).
An exact proof term for the current goal is SNo_add_SNo x u Hx Hu1.
Apply add_SNo_Lt1_cancel (v * z + x * w + y * w) (u * w) (x * z + y * z + v * w) Lvzxwyw Luw Lxzyzvw to the current goal.
We will prove (v * z + x * w + y * w) + u * w < (x * z + y * z + v * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w.
Apply SNoLeLt_tra (v * z + x * w + y * w + u * w) (x * z + u * z + v * w + y * w) (x * z + y * z + v * w + u * w) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) (SNo_add_SNo_4 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) to the current goal.
We will prove v * z + x * w + y * w + u * w x * z + u * z + v * w + y * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove v * z + x * w + u * w + y * w x * z + u * z + v * w + y * w.
rewrite the current goal using add_SNo_rotate_4_1 (v * z) (x * w) (u * w) (y * w) Lvz Lxw Luw Lyw (from left to right).
We will prove y * w + v * z + x * w + u * w x * z + u * z + v * w + y * w.
rewrite the current goal using add_SNo_rotate_4_1 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw (from left to right).
We will prove y * w + v * z + x * w + u * w y * w + x * z + u * z + v * w.
Apply add_SNo_Le2 (y * w) (v * z + x * w + u * w) (x * z + u * z + v * w) Lyw (SNo_add_SNo_3 (v * z) (x * w) (u * w) Lvz Lxw Luw) (SNo_add_SNo_3 (x * z) (u * z) (v * w) Lxz Luz Lvw) to the current goal.
We will prove v * z + x * w + u * w x * z + u * z + v * w.
rewrite the current goal using IHyz u (SNoL_SNoS y Hy u Hu) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove v * z + (x + u) * w x * z + u * z + v * w.
rewrite the current goal using add_SNo_assoc (x * z) (u * z) (v * w) Lxz Luz Lvw (from left to right).
We will prove v * z + (x + u) * w (x * z + u * z) + v * w.
rewrite the current goal using IHy u (SNoL_SNoS y Hy u Hu) (from right to left).
We will prove v * z + (x + u) * w (x + u) * z + v * w.
Apply mul_SNo_Le (x + u) z v w Lxu Hz Hv1 Hw1 to the current goal.
We will prove v x + u.
An exact proof term for the current goal is Hvu.
We will prove w z.
Apply SNoLtLe to the current goal.
We will prove w < z.
An exact proof term for the current goal is Hw3.
We will prove x * z + u * z + v * w + y * w < x * z + y * z + v * w + u * w.
Apply add_SNo_Lt2 (x * z) (u * z + v * w + y * w) (y * z + v * w + u * w) Lxz (SNo_add_SNo_3 (u * z) (v * w) (y * w) Luz Lvw Lyw) (SNo_add_SNo_3 (y * z) (v * w) (u * w) Lyz Lvw Luw) to the current goal.
We will prove u * z + v * w + y * w < y * z + v * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (y * w) Luz Lvw Lyw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * z) (v * w) (u * w) Lyz Lvw Luw (from left to right).
We will prove v * w + u * z + y * w < v * w + y * z + u * w.
Apply add_SNo_Lt2 (v * w) (u * z + y * w) (y * z + u * w) Lvw (SNo_add_SNo (u * z) (y * w) Luz Lyw) (SNo_add_SNo (y * z) (u * w) Lyz Luw) to the current goal.
We will prove u * z + y * w < y * z + u * w.
An exact proof term for the current goal is mul_SNo_Lt y z u w Hy Hz Hu1 Hw1 Hu3 Hw3.
Let v be given.
Assume Hv: v SNoR (x + y).
Let w be given.
Assume Hw: w SNoR z.
Assume Hue: u = v * z + (x + y) * w + - v * w.
rewrite the current goal using Hue (from left to right).
We will prove v * z + (x + y) * w + - v * w < x * z + y * z.
Apply SNoR_E (x + y) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lxwyw: SNo (x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (x * w) (y * w) Lxw Lyw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvzxwyw: SNo (v * z + x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo_3 (v * z) (x * w) (y * w) Lvz Lxw Lyw.
We prove the intermediate claim Lxzyzvw: SNo (x * z + y * z + v * w).
An exact proof term for the current goal is SNo_add_SNo_3 (x * z) (y * z) (v * w) Lxz Lyz Lvw.
Apply add_SNo_minus_Lt1b3 (v * z) ((x + y) * w) (v * w) (x * z + y * z) Lvz Lxyw Lvw Lxzyz to the current goal.
We will prove v * z + (x + y) * w < (x * z + y * z) + v * w.
rewrite the current goal using IHz w (SNoR_SNoS z Hz w Hw) (from left to right).
We will prove v * z + x * w + y * w < (x * z + y * z) + v * w.
rewrite the current goal using add_SNo_assoc (x * z) (y * z) (v * w) Lxz Lyz Lvw (from right to left).
We will prove v * z + x * w + y * w < x * z + y * z + v * w.
Apply add_SNo_SNoR_interpolate x y Hx Hy v Hv to the current goal.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoR x.
Assume Hvu: u + y v.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Luy: SNo (u + y).
An exact proof term for the current goal is SNo_add_SNo u y Hu1 Hy.
Apply add_SNo_Lt1_cancel (v * z + x * w + y * w) (u * w) (x * z + y * z + v * w) Lvzxwyw Luw Lxzyzvw to the current goal.
We will prove (v * z + x * w + y * w) + u * w < (x * z + y * z + v * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w.
Apply SNoLeLt_tra (v * z + x * w + y * w + u * w) (u * z + y * z + v * w + x * w) (x * z + y * z + v * w + u * w) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) (SNo_add_SNo_4 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) to the current goal.
We will prove v * z + x * w + y * w + u * w u * z + y * z + v * w + x * w.
rewrite the current goal using add_SNo_com_3_0_1 (v * z) (x * w) (y * w + u * w) Lvz Lxw (SNo_add_SNo (y * w) (u * w) Lyw Luw) (from left to right).
We will prove x * w + v * z + y * w + u * w u * z + y * z + v * w + x * w.
rewrite the current goal using add_SNo_rotate_4_1 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw (from left to right).
We will prove x * w + v * z + y * w + u * w x * w + u * z + y * z + v * w.
Apply add_SNo_Le2 (x * w) (v * z + y * w + u * w) (u * z + y * z + v * w) Lxw (SNo_add_SNo_3 (v * z) (y * w) (u * w) Lvz Lyw Luw) (SNo_add_SNo_3 (u * z) (y * z) (v * w) Luz Lyz Lvw) to the current goal.
We will prove v * z + y * w + u * w u * z + y * z + v * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove v * z + u * w + y * w u * z + y * z + v * w.
rewrite the current goal using IHxz u (SNoR_SNoS x Hx u Hu) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove v * z + (u + y) * w u * z + y * z + v * w.
rewrite the current goal using add_SNo_assoc (u * z) (y * z) (v * w) Luz Lyz Lvw (from left to right).
We will prove v * z + (u + y) * w (u * z + y * z) + v * w.
rewrite the current goal using IHx u (SNoR_SNoS x Hx u Hu) (from right to left).
We will prove v * z + (u + y) * w (u + y) * z + v * w.
rewrite the current goal using add_SNo_com ((u + y) * z) (v * w) (SNo_mul_SNo (u + y) z Luy Hz) Lvw (from left to right).
rewrite the current goal using add_SNo_com (v * z) ((u + y) * w) Lvz (SNo_mul_SNo (u + y) w Luy Hw1) (from left to right).
Apply mul_SNo_Le v w (u + y) z Hv1 Hw1 Luy Hz to the current goal.
We will prove u + y v.
An exact proof term for the current goal is Hvu.
We will prove z w.
Apply SNoLtLe to the current goal.
We will prove z < w.
An exact proof term for the current goal is Hw3.
We will prove u * z + y * z + v * w + x * w < x * z + y * z + v * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (y * z) (v * w + x * w) Luz Lyz (SNo_add_SNo (v * w) (x * w) Lvw Lxw) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (y * z) (v * w + u * w) Lxz Lyz (SNo_add_SNo (v * w) (u * w) Lvw Luw) (from left to right).
We will prove y * z + u * z + v * w + x * w < y * z + x * z + v * w + u * w.
Apply add_SNo_Lt2 (y * z) (u * z + v * w + x * w) (x * z + v * w + u * w) Lyz (SNo_add_SNo_3 (u * z) (v * w) (x * w) Luz Lvw Lxw) (SNo_add_SNo_3 (x * z) (v * w) (u * w) Lxz Lvw Luw) to the current goal.
We will prove u * z + v * w + x * w < x * z + v * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (x * w) Luz Lvw Lxw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (v * w) (u * w) Lxz Lvw Luw (from left to right).
We will prove v * w + u * z + x * w < v * w + x * z + u * w.
Apply add_SNo_Lt2 (v * w) (u * z + x * w) (x * z + u * w) Lvw (SNo_add_SNo (u * z) (x * w) Luz Lxw) (SNo_add_SNo (x * z) (u * w) Lxz Luw) to the current goal.
We will prove u * z + x * w < x * z + u * w.
rewrite the current goal using add_SNo_com (x * z) (u * w) Lxz Luw (from left to right).
rewrite the current goal using add_SNo_com (u * z) (x * w) Luz Lxw (from left to right).
We will prove x * w + u * z < u * w + x * z.
An exact proof term for the current goal is mul_SNo_Lt u w x z Hu1 Hw1 Hx Hz Hu3 Hw3.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoR y.
Assume Hvu: x + u v.
Apply SNoR_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Lxu: SNo (x + u).
An exact proof term for the current goal is SNo_add_SNo x u Hx Hu1.
Apply add_SNo_Lt1_cancel (v * z + x * w + y * w) (u * w) (x * z + y * z + v * w) Lvzxwyw Luw Lxzyzvw to the current goal.
We will prove (v * z + x * w + y * w) + u * w < (x * z + y * z + v * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w.
Apply SNoLeLt_tra (v * z + x * w + y * w + u * w) (x * z + u * z + v * w + y * w) (x * z + y * z + v * w + u * w) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) (SNo_add_SNo_4 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) to the current goal.
We will prove v * z + x * w + y * w + u * w x * z + u * z + v * w + y * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove v * z + x * w + u * w + y * w x * z + u * z + v * w + y * w.
rewrite the current goal using add_SNo_rotate_4_1 (v * z) (x * w) (u * w) (y * w) Lvz Lxw Luw Lyw (from left to right).
We will prove y * w + v * z + x * w + u * w x * z + u * z + v * w + y * w.
rewrite the current goal using add_SNo_rotate_4_1 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw (from left to right).
We will prove y * w + v * z + x * w + u * w y * w + x * z + u * z + v * w.
Apply add_SNo_Le2 (y * w) (v * z + x * w + u * w) (x * z + u * z + v * w) Lyw (SNo_add_SNo_3 (v * z) (x * w) (u * w) Lvz Lxw Luw) (SNo_add_SNo_3 (x * z) (u * z) (v * w) Lxz Luz Lvw) to the current goal.
We will prove v * z + x * w + u * w x * z + u * z + v * w.
rewrite the current goal using IHyz u (SNoR_SNoS y Hy u Hu) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove v * z + (x + u) * w x * z + u * z + v * w.
rewrite the current goal using add_SNo_assoc (x * z) (u * z) (v * w) Lxz Luz Lvw (from left to right).
We will prove v * z + (x + u) * w (x * z + u * z) + v * w.
rewrite the current goal using IHy u (SNoR_SNoS y Hy u Hu) (from right to left).
We will prove v * z + (x + u) * w (x + u) * z + v * w.
rewrite the current goal using add_SNo_com ((x + u) * z) (v * w) (SNo_mul_SNo (x + u) z Lxu Hz) Lvw (from left to right).
rewrite the current goal using add_SNo_com (v * z) ((x + u) * w) Lvz (SNo_mul_SNo (x + u) w Lxu Hw1) (from left to right).
Apply mul_SNo_Le v w (x + u) z Hv1 Hw1 Lxu Hz to the current goal.
We will prove x + u v.
An exact proof term for the current goal is Hvu.
We will prove z w.
Apply SNoLtLe to the current goal.
We will prove z < w.
An exact proof term for the current goal is Hw3.
We will prove x * z + u * z + v * w + y * w < x * z + y * z + v * w + u * w.
Apply add_SNo_Lt2 (x * z) (u * z + v * w + y * w) (y * z + v * w + u * w) Lxz (SNo_add_SNo_3 (u * z) (v * w) (y * w) Luz Lvw Lyw) (SNo_add_SNo_3 (y * z) (v * w) (u * w) Lyz Lvw Luw) to the current goal.
We will prove u * z + v * w + y * w < y * z + v * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (y * w) Luz Lvw Lyw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * z) (v * w) (u * w) Lyz Lvw Luw (from left to right).
We will prove v * w + u * z + y * w < v * w + y * z + u * w.
Apply add_SNo_Lt2 (v * w) (u * z + y * w) (y * z + u * w) Lvw (SNo_add_SNo (u * z) (y * w) Luz Lyw) (SNo_add_SNo (y * z) (u * w) Lyz Luw) to the current goal.
We will prove u * z + y * w < y * z + u * w.
rewrite the current goal using add_SNo_com (y * z) (u * w) Lyz Luw (from left to right).
rewrite the current goal using add_SNo_com (u * z) (y * w) Luz Lyw (from left to right).
We will prove y * w + u * z < u * w + y * z.
An exact proof term for the current goal is mul_SNo_Lt u w y z Hu1 Hw1 Hy Hz Hu3 Hw3.
Let u be given.
Assume Hu: u R.
rewrite the current goal using LE (from right to left).
We will prove x * z + y * z < u.
Apply HRE u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL (x + y).
Let w be given.
Assume Hw: w SNoR z.
Assume Hue: u = v * z + (x + y) * w + - v * w.
rewrite the current goal using Hue (from left to right).
We will prove x * z + y * z < v * z + (x + y) * w + - v * w.
Apply SNoL_E (x + y) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lxwyw: SNo (x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (x * w) (y * w) Lxw Lyw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvzxwyw: SNo (v * z + x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo_3 (v * z) (x * w) (y * w) Lvz Lxw Lyw.
We prove the intermediate claim Lxzyzvw: SNo (x * z + y * z + v * w).
An exact proof term for the current goal is SNo_add_SNo_3 (x * z) (y * z) (v * w) Lxz Lyz Lvw.
We will prove x * z + y * z < v * z + (x + y) * w + - v * w.
Apply add_SNo_minus_Lt2b3 (v * z) ((x + y) * w) (v * w) (x * z + y * z) Lvz Lxyw Lvw Lxzyz to the current goal.
We will prove (x * z + y * z) + v * w < v * z + (x + y) * w.
rewrite the current goal using IHz w (SNoR_SNoS z Hz w Hw) (from left to right).
We will prove (x * z + y * z) + v * w < v * z + x * w + y * w.
rewrite the current goal using add_SNo_assoc (x * z) (y * z) (v * w) Lxz Lyz Lvw (from right to left).
We will prove x * z + y * z + v * w < v * z + x * w + y * w.
Apply add_SNo_SNoL_interpolate x y Hx Hy v Hv to the current goal.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoL x.
Assume Hvu: v u + y.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Luy: SNo (u + y).
An exact proof term for the current goal is SNo_add_SNo u y Hu1 Hy.
Apply add_SNo_Lt1_cancel (x * z + y * z + v * w) (u * w) (v * z + x * w + y * w) Lxzyzvw Luw Lvzxwyw to the current goal.
We will prove (x * z + y * z + v * w) + u * w < (v * z + x * w + y * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w.
Apply SNoLtLe_tra (x * z + y * z + v * w + u * w) (u * z + y * z + v * w + x * w) (v * z + x * w + y * w + u * w) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) (SNo_add_SNo_4 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) to the current goal.
We will prove x * z + y * z + v * w + u * w < u * z + y * z + v * w + x * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (y * z) (v * w + x * w) Luz Lyz (SNo_add_SNo (v * w) (x * w) Lvw Lxw) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (y * z) (v * w + u * w) Lxz Lyz (SNo_add_SNo (v * w) (u * w) Lvw Luw) (from left to right).
We will prove y * z + x * z + v * w + u * w < y * z + u * z + v * w + x * w.
Apply add_SNo_Lt2 (y * z) (x * z + v * w + u * w) (u * z + v * w + x * w) Lyz (SNo_add_SNo_3 (x * z) (v * w) (u * w) Lxz Lvw Luw) (SNo_add_SNo_3 (u * z) (v * w) (x * w) Luz Lvw Lxw) to the current goal.
We will prove x * z + v * w + u * w < u * z + v * w + x * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (x * w) Luz Lvw Lxw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (v * w) (u * w) Lxz Lvw Luw (from left to right).
We will prove v * w + x * z + u * w < v * w + u * z + x * w.
Apply add_SNo_Lt2 (v * w) (x * z + u * w) (u * z + x * w) Lvw (SNo_add_SNo (x * z) (u * w) Lxz Luw) (SNo_add_SNo (u * z) (x * w) Luz Lxw) to the current goal.
We will prove x * z + u * w < u * z + x * w.
rewrite the current goal using add_SNo_com (x * z) (u * w) Lxz Luw (from left to right).
rewrite the current goal using add_SNo_com (u * z) (x * w) Luz Lxw (from left to right).
An exact proof term for the current goal is mul_SNo_Lt x w u z Hx Hw1 Hu1 Hz Hu3 Hw3.
We will prove u * z + y * z + v * w + x * w v * z + x * w + y * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (v * z) (x * w) (y * w + u * w) Lvz Lxw (SNo_add_SNo (y * w) (u * w) Lyw Luw) (from left to right).
We will prove u * z + y * z + v * w + x * w x * w + v * z + y * w + u * w.
rewrite the current goal using add_SNo_rotate_4_1 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw (from left to right).
We will prove x * w + u * z + y * z + v * w x * w + v * z + y * w + u * w.
Apply add_SNo_Le2 (x * w) (u * z + y * z + v * w) (v * z + y * w + u * w) Lxw (SNo_add_SNo_3 (u * z) (y * z) (v * w) Luz Lyz Lvw) (SNo_add_SNo_3 (v * z) (y * w) (u * w) Lvz Lyw Luw) to the current goal.
We will prove u * z + y * z + v * w v * z + y * w + u * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove u * z + y * z + v * w v * z + u * w + y * w.
rewrite the current goal using IHxz u (SNoL_SNoS x Hx u Hu) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove u * z + y * z + v * w v * z + (u + y) * w.
rewrite the current goal using add_SNo_assoc (u * z) (y * z) (v * w) Luz Lyz Lvw (from left to right).
We will prove (u * z + y * z) + v * w v * z + (u + y) * w.
rewrite the current goal using IHx u (SNoL_SNoS x Hx u Hu) (from right to left).
We will prove (u + y) * z + v * w v * z + (u + y) * w.
rewrite the current goal using add_SNo_com ((u + y) * z) (v * w) (SNo_mul_SNo (u + y) z Luy Hz) Lvw (from left to right).
rewrite the current goal using add_SNo_com (v * z) ((u + y) * w) Lvz (SNo_mul_SNo (u + y) w Luy Hw1) (from left to right).
We will prove v * w + (u + y) * z (u + y) * w + v * z.
Apply mul_SNo_Le (u + y) w v z Luy Hw1 Hv1 Hz to the current goal.
We will prove v u + y.
An exact proof term for the current goal is Hvu.
We will prove z w.
Apply SNoLtLe to the current goal.
We will prove z < w.
An exact proof term for the current goal is Hw3.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoL y.
Assume Hvu: v x + u.
Apply SNoL_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Lxu: SNo (x + u).
An exact proof term for the current goal is SNo_add_SNo x u Hx Hu1.
Apply add_SNo_Lt1_cancel (x * z + y * z + v * w) (u * w) (v * z + x * w + y * w) Lxzyzvw Luw Lvzxwyw to the current goal.
We will prove (x * z + y * z + v * w) + u * w < (v * z + x * w + y * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w.
Apply SNoLtLe_tra (x * z + y * z + v * w + u * w) (x * z + u * z + v * w + y * w) (v * z + x * w + y * w + u * w) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) (SNo_add_SNo_4 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) to the current goal.
We will prove x * z + y * z + v * w + u * w < x * z + u * z + v * w + y * w.
Apply add_SNo_Lt2 (x * z) (y * z + v * w + u * w) (u * z + v * w + y * w) Lxz (SNo_add_SNo_3 (y * z) (v * w) (u * w) Lyz Lvw Luw) (SNo_add_SNo_3 (u * z) (v * w) (y * w) Luz Lvw Lyw) to the current goal.
We will prove y * z + v * w + u * w < u * z + v * w + y * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (y * w) Luz Lvw Lyw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * z) (v * w) (u * w) Lyz Lvw Luw (from left to right).
We will prove v * w + y * z + u * w < v * w + u * z + y * w.
Apply add_SNo_Lt2 (v * w) (y * z + u * w) (u * z + y * w) Lvw (SNo_add_SNo (y * z) (u * w) Lyz Luw) (SNo_add_SNo (u * z) (y * w) Luz Lyw) to the current goal.
We will prove y * z + u * w < u * z + y * w.
rewrite the current goal using add_SNo_com (y * z) (u * w) Lyz Luw (from left to right).
rewrite the current goal using add_SNo_com (u * z) (y * w) Luz Lyw (from left to right).
An exact proof term for the current goal is mul_SNo_Lt y w u z Hy Hw1 Hu1 Hz Hu3 Hw3.
We will prove x * z + u * z + v * w + y * w v * z + x * w + y * w + u * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove x * z + u * z + v * w + y * w v * z + x * w + u * w + y * w.
rewrite the current goal using add_SNo_rotate_4_1 (v * z) (x * w) (u * w) (y * w) Lvz Lxw Luw Lyw (from left to right).
We will prove x * z + u * z + v * w + y * w y * w + v * z + x * w + u * w.
rewrite the current goal using add_SNo_rotate_4_1 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw (from left to right).
We will prove y * w + x * z + u * z + v * w y * w + v * z + x * w + u * w.
Apply add_SNo_Le2 (y * w) (x * z + u * z + v * w) (v * z + x * w + u * w) Lyw (SNo_add_SNo_3 (x * z) (u * z) (v * w) Lxz Luz Lvw) (SNo_add_SNo_3 (v * z) (x * w) (u * w) Lvz Lxw Luw) to the current goal.
We will prove x * z + u * z + v * w v * z + x * w + u * w.
rewrite the current goal using IHyz u (SNoL_SNoS y Hy u Hu) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove x * z + u * z + v * w v * z + (x + u) * w.
rewrite the current goal using add_SNo_assoc (x * z) (u * z) (v * w) Lxz Luz Lvw (from left to right).
We will prove (x * z + u * z) + v * w v * z + (x + u) * w.
rewrite the current goal using IHy u (SNoL_SNoS y Hy u Hu) (from right to left).
We will prove (x + u) * z + v * w v * z + (x + u) * w.
rewrite the current goal using add_SNo_com ((x + u) * z) (v * w) (SNo_mul_SNo (x + u) z Lxu Hz) Lvw (from left to right).
rewrite the current goal using add_SNo_com (v * z) ((x + u) * w) Lvz (SNo_mul_SNo (x + u) w Lxu Hw1) (from left to right).
We will prove v * w + (x + u) * z (x + u) * w + v * z.
Apply mul_SNo_Le (x + u) w v z Lxu Hw1 Hv1 Hz to the current goal.
We will prove v x + u.
An exact proof term for the current goal is Hvu.
We will prove z w.
Apply SNoLtLe to the current goal.
We will prove z < w.
An exact proof term for the current goal is Hw3.
Let v be given.
Assume Hv: v SNoR (x + y).
Let w be given.
Assume Hw: w SNoL z.
Assume Hue: u = v * z + (x + y) * w + - v * w.
rewrite the current goal using Hue (from left to right).
We will prove x * z + y * z < v * z + (x + y) * w + - v * w.
Apply SNoR_E (x + y) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lxwyw: SNo (x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (x * w) (y * w) Lxw Lyw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvzxwyw: SNo (v * z + x * w + y * w).
An exact proof term for the current goal is SNo_add_SNo_3 (v * z) (x * w) (y * w) Lvz Lxw Lyw.
We prove the intermediate claim Lxzyzvw: SNo (x * z + y * z + v * w).
An exact proof term for the current goal is SNo_add_SNo_3 (x * z) (y * z) (v * w) Lxz Lyz Lvw.
We will prove x * z + y * z < v * z + (x + y) * w + - v * w.
Apply add_SNo_minus_Lt2b3 (v * z) ((x + y) * w) (v * w) (x * z + y * z) Lvz Lxyw Lvw Lxzyz to the current goal.
We will prove (x * z + y * z) + v * w < v * z + (x + y) * w.
rewrite the current goal using IHz w (SNoL_SNoS z Hz w Hw) (from left to right).
We will prove (x * z + y * z) + v * w < v * z + x * w + y * w.
rewrite the current goal using add_SNo_assoc (x * z) (y * z) (v * w) Lxz Lyz Lvw (from right to left).
We will prove x * z + y * z + v * w < v * z + x * w + y * w.
Apply add_SNo_SNoR_interpolate x y Hx Hy v Hv to the current goal.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoR x.
Assume Hvu: u + y v.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Luy: SNo (u + y).
An exact proof term for the current goal is SNo_add_SNo u y Hu1 Hy.
Apply add_SNo_Lt1_cancel (x * z + y * z + v * w) (u * w) (v * z + x * w + y * w) Lxzyzvw Luw Lvzxwyw to the current goal.
We will prove (x * z + y * z + v * w) + u * w < (v * z + x * w + y * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w.
Apply SNoLtLe_tra (x * z + y * z + v * w + u * w) (u * z + y * z + v * w + x * w) (v * z + x * w + y * w + u * w) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) (SNo_add_SNo_4 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) to the current goal.
We will prove x * z + y * z + v * w + u * w < u * z + y * z + v * w + x * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (y * z) (v * w + x * w) Luz Lyz (SNo_add_SNo (v * w) (x * w) Lvw Lxw) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (y * z) (v * w + u * w) Lxz Lyz (SNo_add_SNo (v * w) (u * w) Lvw Luw) (from left to right).
We will prove y * z + x * z + v * w + u * w < y * z + u * z + v * w + x * w.
Apply add_SNo_Lt2 (y * z) (x * z + v * w + u * w) (u * z + v * w + x * w) Lyz (SNo_add_SNo_3 (x * z) (v * w) (u * w) Lxz Lvw Luw) (SNo_add_SNo_3 (u * z) (v * w) (x * w) Luz Lvw Lxw) to the current goal.
We will prove x * z + v * w + u * w < u * z + v * w + x * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (x * w) Luz Lvw Lxw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (x * z) (v * w) (u * w) Lxz Lvw Luw (from left to right).
We will prove v * w + x * z + u * w < v * w + u * z + x * w.
Apply add_SNo_Lt2 (v * w) (x * z + u * w) (u * z + x * w) Lvw (SNo_add_SNo (x * z) (u * w) Lxz Luw) (SNo_add_SNo (u * z) (x * w) Luz Lxw) to the current goal.
We will prove x * z + u * w < u * z + x * w.
An exact proof term for the current goal is mul_SNo_Lt u z x w Hu1 Hz Hx Hw1 Hu3 Hw3.
We will prove u * z + y * z + v * w + x * w v * z + x * w + y * w + u * w.
rewrite the current goal using add_SNo_com_3_0_1 (v * z) (x * w) (y * w + u * w) Lvz Lxw (SNo_add_SNo (y * w) (u * w) Lyw Luw) (from left to right).
We will prove u * z + y * z + v * w + x * w x * w + v * z + y * w + u * w.
rewrite the current goal using add_SNo_rotate_4_1 (u * z) (y * z) (v * w) (x * w) Luz Lyz Lvw Lxw (from left to right).
We will prove x * w + u * z + y * z + v * w x * w + v * z + y * w + u * w.
Apply add_SNo_Le2 (x * w) (u * z + y * z + v * w) (v * z + y * w + u * w) Lxw (SNo_add_SNo_3 (u * z) (y * z) (v * w) Luz Lyz Lvw) (SNo_add_SNo_3 (v * z) (y * w) (u * w) Lvz Lyw Luw) to the current goal.
We will prove u * z + y * z + v * w v * z + y * w + u * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove u * z + y * z + v * w v * z + u * w + y * w.
rewrite the current goal using IHxz u (SNoR_SNoS x Hx u Hu) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove u * z + y * z + v * w v * z + (u + y) * w.
rewrite the current goal using add_SNo_assoc (u * z) (y * z) (v * w) Luz Lyz Lvw (from left to right).
We will prove (u * z + y * z) + v * w v * z + (u + y) * w.
rewrite the current goal using IHx u (SNoR_SNoS x Hx u Hu) (from right to left).
We will prove (u + y) * z + v * w v * z + (u + y) * w.
Apply mul_SNo_Le v z (u + y) w Hv1 Hz Luy Hw1 to the current goal.
We will prove u + y v.
An exact proof term for the current goal is Hvu.
We will prove w z.
Apply SNoLtLe to the current goal.
We will prove w < z.
An exact proof term for the current goal is Hw3.
Assume H1.
Apply H1 to the current goal.
Let u be given.
Assume H1.
Apply H1 to the current goal.
Assume Hu: u SNoR y.
Assume Hvu: x + u v.
Apply SNoR_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
We prove the intermediate claim Luw: SNo (u * w).
An exact proof term for the current goal is SNo_mul_SNo u w Hu1 Hw1.
We prove the intermediate claim Luz: SNo (u * z).
An exact proof term for the current goal is SNo_mul_SNo u z Hu1 Hz.
We prove the intermediate claim Lxu: SNo (x + u).
An exact proof term for the current goal is SNo_add_SNo x u Hx Hu1.
Apply add_SNo_Lt1_cancel (x * z + y * z + v * w) (u * w) (v * z + x * w + y * w) Lxzyzvw Luw Lvzxwyw to the current goal.
We will prove (x * z + y * z + v * w) + u * w < (v * z + x * w + y * w) + u * w.
rewrite the current goal using add_SNo_assoc_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using add_SNo_assoc_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw (from right to left).
We will prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w.
Apply SNoLtLe_tra (x * z + y * z + v * w + u * w) (x * z + u * z + v * w + y * w) (v * z + x * w + y * w + u * w) (SNo_add_SNo_4 (x * z) (y * z) (v * w) (u * w) Lxz Lyz Lvw Luw) (SNo_add_SNo_4 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw) (SNo_add_SNo_4 (v * z) (x * w) (y * w) (u * w) Lvz Lxw Lyw Luw) to the current goal.
We will prove x * z + y * z + v * w + u * w < x * z + u * z + v * w + y * w.
Apply add_SNo_Lt2 (x * z) (y * z + v * w + u * w) (u * z + v * w + y * w) Lxz (SNo_add_SNo_3 (y * z) (v * w) (u * w) Lyz Lvw Luw) (SNo_add_SNo_3 (u * z) (v * w) (y * w) Luz Lvw Lyw) to the current goal.
We will prove y * z + v * w + u * w < u * z + v * w + y * w.
rewrite the current goal using add_SNo_com_3_0_1 (u * z) (v * w) (y * w) Luz Lvw Lyw (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (y * z) (v * w) (u * w) Lyz Lvw Luw (from left to right).
We will prove v * w + y * z + u * w < v * w + u * z + y * w.
Apply add_SNo_Lt2 (v * w) (y * z + u * w) (u * z + y * w) Lvw (SNo_add_SNo (y * z) (u * w) Lyz Luw) (SNo_add_SNo (u * z) (y * w) Luz Lyw) to the current goal.
We will prove y * z + u * w < u * z + y * w.
An exact proof term for the current goal is mul_SNo_Lt u z y w Hu1 Hz Hy Hw1 Hu3 Hw3.
We will prove x * z + u * z + v * w + y * w v * z + x * w + y * w + u * w.
rewrite the current goal using add_SNo_com (y * w) (u * w) Lyw Luw (from left to right).
We will prove x * z + u * z + v * w + y * w v * z + x * w + u * w + y * w.
rewrite the current goal using add_SNo_rotate_4_1 (v * z) (x * w) (u * w) (y * w) Lvz Lxw Luw Lyw (from left to right).
We will prove x * z + u * z + v * w + y * w y * w + v * z + x * w + u * w.
rewrite the current goal using add_SNo_rotate_4_1 (x * z) (u * z) (v * w) (y * w) Lxz Luz Lvw Lyw (from left to right).
We will prove y * w + x * z + u * z + v * w y * w + v * z + x * w + u * w.
Apply add_SNo_Le2 (y * w) (x * z + u * z + v * w) (v * z + x * w + u * w) Lyw (SNo_add_SNo_3 (x * z) (u * z) (v * w) Lxz Luz Lvw) (SNo_add_SNo_3 (v * z) (x * w) (u * w) Lvz Lxw Luw) to the current goal.
We will prove x * z + u * z + v * w v * z + x * w + u * w.
rewrite the current goal using IHyz u (SNoR_SNoS y Hy u Hu) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove x * z + u * z + v * w v * z + (x + u) * w.
rewrite the current goal using add_SNo_assoc (x * z) (u * z) (v * w) Lxz Luz Lvw (from left to right).
We will prove (x * z + u * z) + v * w v * z + (x + u) * w.
rewrite the current goal using IHy u (SNoR_SNoS y Hy u Hu) (from right to left).
We will prove (x + u) * z + v * w v * z + (x + u) * w.
Apply mul_SNo_Le v z (x + u) w Hv1 Hz Lxu Hw1 to the current goal.
We will prove x + u v.
An exact proof term for the current goal is Hvu.
We will prove w z.
Apply SNoLtLe to the current goal.
We will prove w < z.
An exact proof term for the current goal is Hw3.
Let u' be given.
Assume Hu': u' L1 L2.
rewrite the current goal using HE (from right to left).
We will prove u' < (x + y) * z.
Apply binunionE L1 L2 u' Hu' to the current goal.
Assume Hu': u' L1.
Apply ReplE_impred (SNoL (x * z)) (λw ⇒ w + y * z) u' Hu' to the current goal.
Let u be given.
Assume Hu: u SNoL (x * z).
Assume Hu'u: u' = u + y * z.
rewrite the current goal using Hu'u (from left to right).
We will prove u + y * z < (x + y) * z.
Apply SNoL_E (x * z) Lxz u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
Apply mul_SNo_SNoL_interpolate_impred x z Hx Hz u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoL z.
Assume Hvw: u + v * w v * z + x * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvy: SNo (v + y).
An exact proof term for the current goal is SNo_add_SNo v y Hv1 Hy.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luyz: SNo (u + y * z).
An exact proof term for the current goal is SNo_add_SNo u (y * z) Hu1 Lyz.
We prove the intermediate claim Lvwyw: SNo (v * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (y * w) Lvw Lyw.
We prove the intermediate claim Lvzxw: SNo (v * z + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * z) (x * w) Lvz Lxw.
We will prove u + y * z < (x + y) * z.
Apply add_SNo_Lt1_cancel (u + y * z) (v * w + y * w) ((x + y) * z) Luyz Lvwyw Lxyz to the current goal.
We will prove (u + y * z) + v * w + y * w < (x + y) * z + v * w + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (y * z) (v * w) (y * w) Hu1 Lyz Lvw Lyw (from left to right).
We will prove (u + v * w) + y * z + y * w < (x + y) * z + v * w + y * w.
Apply SNoLeLt_tra ((u + v * w) + y * z + y * w) (v * z + y * z + x * w + y * w) ((x + y) * z + v * w + y * w) (SNo_add_SNo_3 (u + v * w) (y * z) (y * w) Luvw Lyz Lyw) (SNo_add_SNo_4 (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw) (SNo_add_SNo_3 ((x + y) * z) (v * w) (y * w) Lxyz Lvw Lyw) to the current goal.
We will prove (u + v * w) + y * z + y * w v * z + y * z + x * w + y * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) (x * w + y * w) Lvz Lyz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
We will prove (u + v * w) + y * z + y * w (v * z + y * z) + x * w + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw (from left to right).
We will prove (u + v * w) + y * z + y * w (v * z + x * w) + y * z + y * w.
Apply add_SNo_Le1 (u + v * w) (y * z + y * w) (v * z + x * w) Luvw (SNo_add_SNo (y * z) (y * w) Lyz Lyw) (SNo_add_SNo (v * z) (x * w) Lvz Lxw) to the current goal.
We will prove u + v * w v * z + x * w.
An exact proof term for the current goal is Hvw.
We will prove v * z + y * z + x * w + y * w < (x + y) * z + v * w + y * w.
rewrite the current goal using IHz w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove v * z + y * z + (x + y) * w < (x + y) * z + v * w + y * w.
rewrite the current goal using IHxz v (SNoL_SNoS x Hx v Hv) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove v * z + y * z + (x + y) * w < (x + y) * z + (v + y) * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) ((x + y) * w) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using IHx v (SNoL_SNoS x Hx v Hv) (from right to left).
We will prove (v + y) * z + (x + y) * w < (x + y) * z + (v + y) * w.
Apply mul_SNo_Lt (x + y) z (v + y) w Lxy Hz Lvy Hw1 to the current goal.
We will prove v + y < x + y.
Apply add_SNo_Lt1 v y x Hv1 Hy Hx to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w < z.
An exact proof term for the current goal is Hw3.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoR z.
Assume Hvw: u + v * w v * z + x * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvy: SNo (v + y).
An exact proof term for the current goal is SNo_add_SNo v y Hv1 Hy.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luyz: SNo (u + y * z).
An exact proof term for the current goal is SNo_add_SNo u (y * z) Hu1 Lyz.
We prove the intermediate claim Lvwyw: SNo (v * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (y * w) Lvw Lyw.
We will prove u + y * z < (x + y) * z.
Apply add_SNo_Lt1_cancel (u + y * z) (v * w + y * w) ((x + y) * z) Luyz Lvwyw Lxyz to the current goal.
We will prove (u + y * z) + v * w + y * w < (x + y) * z + v * w + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (y * z) (v * w) (y * w) Hu1 Lyz Lvw Lyw (from left to right).
We will prove (u + v * w) + y * z + y * w < (x + y) * z + v * w + y * w.
Apply SNoLeLt_tra ((u + v * w) + y * z + y * w) (v * z + y * z + x * w + y * w) ((x + y) * z + v * w + y * w) (SNo_add_SNo_3 (u + v * w) (y * z) (y * w) Luvw Lyz Lyw) (SNo_add_SNo_4 (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw) (SNo_add_SNo_3 ((x + y) * z) (v * w) (y * w) Lxyz Lvw Lyw) to the current goal.
We will prove (u + v * w) + y * z + y * w v * z + y * z + x * w + y * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) (x * w + y * w) Lvz Lyz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
We will prove (u + v * w) + y * z + y * w (v * z + y * z) + x * w + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw (from left to right).
We will prove (u + v * w) + y * z + y * w (v * z + x * w) + y * z + y * w.
Apply add_SNo_Le1 (u + v * w) (y * z + y * w) (v * z + x * w) Luvw (SNo_add_SNo (y * z) (y * w) Lyz Lyw) (SNo_add_SNo (v * z) (x * w) Lvz Lxw) to the current goal.
We will prove u + v * w v * z + x * w.
An exact proof term for the current goal is Hvw.
We will prove v * z + y * z + x * w + y * w < (x + y) * z + v * w + y * w.
rewrite the current goal using IHz w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove v * z + y * z + (x + y) * w < (x + y) * z + v * w + y * w.
rewrite the current goal using IHxz v (SNoR_SNoS x Hx v Hv) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove v * z + y * z + (x + y) * w < (x + y) * z + (v + y) * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) ((x + y) * w) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using IHx v (SNoR_SNoS x Hx v Hv) (from right to left).
We will prove (v + y) * z + (x + y) * w < (x + y) * z + (v + y) * w.
rewrite the current goal using add_SNo_com ((v + y) * z) ((x + y) * w) (SNo_mul_SNo (v + y) z Lvy Hz) Lxyw (from left to right).
rewrite the current goal using add_SNo_com ((x + y) * z) ((v + y) * w) Lxyz (SNo_mul_SNo (v + y) w Lvy Hw1) (from left to right).
Apply mul_SNo_Lt (v + y) w (x + y) z Lvy Hw1 Lxy Hz to the current goal.
We will prove x + y < v + y.
Apply add_SNo_Lt1 x y v Hx Hy Hv1 to the current goal.
An exact proof term for the current goal is Hv3.
We will prove z < w.
An exact proof term for the current goal is Hw3.
Assume Hu': u' L2.
Apply ReplE_impred (SNoL (y * z)) (λw ⇒ x * z + w) u' Hu' to the current goal.
Let u be given.
Assume Hu: u SNoL (y * z).
Assume Hu'u: u' = x * z + u.
rewrite the current goal using Hu'u (from left to right).
We will prove x * z + u < (x + y) * z.
Apply SNoL_E (y * z) Lyz u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
rewrite the current goal using add_SNo_com (x * z) u Lxz Hu1 (from left to right).
We will prove u + x * z < (x + y) * z.
Apply mul_SNo_SNoL_interpolate_impred y z Hy Hz u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL y.
Let w be given.
Assume Hw: w SNoL z.
Assume Hvw: u + v * w v * z + y * w.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lxv: SNo (x + v).
An exact proof term for the current goal is SNo_add_SNo x v Hx Hv1.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luxz: SNo (u + x * z).
An exact proof term for the current goal is SNo_add_SNo u (x * z) Hu1 Lxz.
We prove the intermediate claim Lvwxw: SNo (v * w + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (x * w) Lvw Lxw.
We will prove u + x * z < (x + y) * z.
Apply add_SNo_Lt1_cancel (u + x * z) (v * w + x * w) ((x + y) * z) Luxz Lvwxw Lxyz to the current goal.
We will prove (u + x * z) + v * w + x * w < (x + y) * z + v * w + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (x * z) (v * w) (x * w) Hu1 Lxz Lvw Lxw (from left to right).
We will prove (u + v * w) + x * z + x * w < (x + y) * z + v * w + x * w.
Apply SNoLeLt_tra ((u + v * w) + x * z + x * w) (v * z + x * z + x * w + y * w) ((x + y) * z + v * w + x * w) (SNo_add_SNo_3 (u + v * w) (x * z) (x * w) Luvw Lxz Lxw) (SNo_add_SNo_4 (v * z) (x * z) (x * w) (y * w) Lvz Lxz Lxw Lyw) (SNo_add_SNo_3 ((x + y) * z) (v * w) (x * w) Lxyz Lvw Lxw) to the current goal.
We will prove (u + v * w) + x * z + x * w v * z + x * z + x * w + y * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) (x * w + y * w) Lvz Lxz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
We will prove (u + v * w) + x * z + x * w (v * z + x * z) + x * w + y * w.
rewrite the current goal using add_SNo_com (x * w) (y * w) Lxw Lyw (from left to right).
We will prove (u + v * w) + x * z + x * w (v * z + x * z) + y * w + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (x * z) (y * w) (x * w) Lvz Lxz Lyw Lxw (from left to right).
We will prove (u + v * w) + x * z + x * w (v * z + y * w) + x * z + x * w.
Apply add_SNo_Le1 (u + v * w) (x * z + x * w) (v * z + y * w) Luvw (SNo_add_SNo (x * z) (x * w) Lxz Lxw) (SNo_add_SNo (v * z) (y * w) Lvz Lyw) to the current goal.
We will prove u + v * w v * z + y * w.
An exact proof term for the current goal is Hvw.
We will prove v * z + x * z + x * w + y * w < (x + y) * z + v * w + x * w.
rewrite the current goal using IHz w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove v * z + x * z + (x + y) * w < (x + y) * z + v * w + x * w.
rewrite the current goal using add_SNo_com (v * w) (x * w) Lvw Lxw (from left to right).
rewrite the current goal using IHyz v (SNoL_SNoS y Hy v Hv) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove v * z + x * z + (x + y) * w < (x + y) * z + (x + v) * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) ((x + y) * w) Lvz Lxz Lxyw (from left to right).
We will prove (v * z + x * z) + (x + y) * w < (x + y) * z + (x + v) * w.
rewrite the current goal using add_SNo_com (v * z) (x * z) Lvz Lxz (from left to right).
We will prove (x * z + v * z) + (x + y) * w < (x + y) * z + (x + v) * w.
rewrite the current goal using IHy v (SNoL_SNoS y Hy v Hv) (from right to left).
We will prove (x + v) * z + (x + y) * w < (x + y) * z + (x + v) * w.
Apply mul_SNo_Lt (x + y) z (x + v) w Lxy Hz Lxv Hw1 to the current goal.
We will prove x + v < x + y.
Apply add_SNo_Lt2 x v y Hx Hv1 Hy to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w < z.
An exact proof term for the current goal is Hw3.
Let v be given.
Assume Hv: v SNoR y.
Let w be given.
Assume Hw: w SNoR z.
Assume Hvw: u + v * w v * z + y * w.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lxv: SNo (x + v).
An exact proof term for the current goal is SNo_add_SNo x v Hx Hv1.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luxz: SNo (u + x * z).
An exact proof term for the current goal is SNo_add_SNo u (x * z) Hu1 Lxz.
We prove the intermediate claim Lvwxw: SNo (v * w + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (x * w) Lvw Lxw.
We prove the intermediate claim Lvzxw: SNo (v * z + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * z) (x * w) Lvz Lxw.
We will prove u + x * z < (x + y) * z.
Apply add_SNo_Lt1_cancel (u + x * z) (v * w + x * w) ((x + y) * z) Luxz Lvwxw Lxyz to the current goal.
We will prove (u + x * z) + v * w + x * w < (x + y) * z + v * w + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (x * z) (v * w) (x * w) Hu1 Lxz Lvw Lxw (from left to right).
We will prove (u + v * w) + x * z + x * w < (x + y) * z + v * w + x * w.
Apply SNoLeLt_tra ((u + v * w) + x * z + x * w) (v * z + x * z + x * w + y * w) ((x + y) * z + v * w + x * w) (SNo_add_SNo_3 (u + v * w) (x * z) (x * w) Luvw Lxz Lxw) (SNo_add_SNo_4 (v * z) (x * z) (x * w) (y * w) Lvz Lxz Lxw Lyw) (SNo_add_SNo_3 ((x + y) * z) (v * w) (x * w) Lxyz Lvw Lxw) to the current goal.
We will prove (u + v * w) + x * z + x * w v * z + x * z + x * w + y * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) (x * w + y * w) Lvz Lxz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
rewrite the current goal using add_SNo_com (y * w) (x * w) Lyw Lxw (from right to left).
We will prove (u + v * w) + x * z + x * w (v * z + x * z) + y * w + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (x * z) (y * w) (x * w) Lvz Lxz Lyw Lxw (from left to right).
We will prove (u + v * w) + x * z + x * w (v * z + y * w) + x * z + x * w.
Apply add_SNo_Le1 (u + v * w) (x * z + x * w) (v * z + y * w) Luvw (SNo_add_SNo (x * z) (x * w) Lxz Lxw) (SNo_add_SNo (v * z) (y * w) Lvz Lyw) to the current goal.
We will prove u + v * w v * z + y * w.
An exact proof term for the current goal is Hvw.
We will prove v * z + x * z + x * w + y * w < (x + y) * z + v * w + x * w.
rewrite the current goal using IHz w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove v * z + x * z + (x + y) * w < (x + y) * z + v * w + x * w.
rewrite the current goal using add_SNo_com (v * w) (x * w) Lvw Lxw (from left to right).
We will prove v * z + x * z + (x + y) * w < (x + y) * z + x * w + v * w.
rewrite the current goal using IHyz v (SNoR_SNoS y Hy v Hv) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove v * z + x * z + (x + y) * w < (x + y) * z + (x + v) * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) ((x + y) * w) Lvz Lxz Lxyw (from left to right).
We will prove (v * z + x * z) + (x + y) * w < (x + y) * z + (x + v) * w.
rewrite the current goal using add_SNo_com (v * z) (x * z) Lvz Lxz (from left to right).
rewrite the current goal using IHy v (SNoR_SNoS y Hy v Hv) (from right to left).
We will prove (x + v) * z + (x + y) * w < (x + y) * z + (x + v) * w.
rewrite the current goal using add_SNo_com ((x + v) * z) ((x + y) * w) (SNo_mul_SNo (x + v) z Lxv Hz) Lxyw (from left to right).
rewrite the current goal using add_SNo_com ((x + y) * z) ((x + v) * w) Lxyz (SNo_mul_SNo (x + v) w Lxv Hw1) (from left to right).
We will prove (x + y) * w + (x + v) * z < (x + v) * w + (x + y) * z.
Apply mul_SNo_Lt (x + v) w (x + y) z Lxv Hw1 Lxy Hz to the current goal.
We will prove x + y < x + v.
Apply add_SNo_Lt2 x y v Hx Hy Hv1 to the current goal.
An exact proof term for the current goal is Hv3.
We will prove z < w.
An exact proof term for the current goal is Hw3.
Let u' be given.
Assume Hu': u' R1 R2.
rewrite the current goal using HE (from right to left).
We will prove (x + y) * z < u'.
Apply binunionE R1 R2 u' Hu' to the current goal.
Assume Hu': u' R1.
Apply ReplE_impred (SNoR (x * z)) (λw ⇒ w + y * z) u' Hu' to the current goal.
Let u be given.
Assume Hu: u SNoR (x * z).
Assume Hu'u: u' = u + y * z.
rewrite the current goal using Hu'u (from left to right).
We will prove (x + y) * z < u + y * z.
Apply SNoR_E (x * z) Lxz u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
Apply mul_SNo_SNoR_interpolate_impred x z Hx Hz u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoR z.
Assume Hvw: v * z + x * w u + v * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvy: SNo (v + y).
An exact proof term for the current goal is SNo_add_SNo v y Hv1 Hy.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luyz: SNo (u + y * z).
An exact proof term for the current goal is SNo_add_SNo u (y * z) Hu1 Lyz.
We prove the intermediate claim Lvwyw: SNo (v * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (y * w) Lvw Lyw.
We prove the intermediate claim Lvzxw: SNo (v * z + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * z) (x * w) Lvz Lxw.
We will prove (x + y) * z < u + y * z.
Apply add_SNo_Lt1_cancel ((x + y) * z) (v * w + y * w) (u + y * z) Lxyz Lvwyw Luyz to the current goal.
We will prove (x + y) * z + v * w + y * w < (u + y * z) + v * w + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (y * z) (v * w) (y * w) Hu1 Lyz Lvw Lyw (from left to right).
We will prove (x + y) * z + v * w + y * w < (u + v * w) + y * z + y * w.
Apply SNoLtLe_tra ((x + y) * z + v * w + y * w) (v * z + y * z + x * w + y * w) ((u + v * w) + y * z + y * w) (SNo_add_SNo_3 ((x + y) * z) (v * w) (y * w) Lxyz Lvw Lyw) (SNo_add_SNo_4 (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw) (SNo_add_SNo_3 (u + v * w) (y * z) (y * w) Luvw Lyz Lyw) to the current goal.
We will prove (x + y) * z + v * w + y * w < v * z + y * z + x * w + y * w.
rewrite the current goal using IHz w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + v * w + y * w < v * z + y * z + (x + y) * w.
rewrite the current goal using IHxz v (SNoL_SNoS x Hx v Hv) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + (v + y) * w < v * z + y * z + (x + y) * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) ((x + y) * w) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using IHx v (SNoL_SNoS x Hx v Hv) (from right to left).
We will prove (x + y) * z + (v + y) * w < (v + y) * z + (x + y) * w.
rewrite the current goal using add_SNo_com ((x + y) * z) ((v + y) * w) Lxyz (SNo_mul_SNo (v + y) w Lvy Hw1) (from left to right).
rewrite the current goal using add_SNo_com ((v + y) * z) ((x + y) * w) (SNo_mul_SNo (v + y) z Lvy Hz) Lxyw (from left to right).
We will prove (v + y) * w + (x + y) * z < (x + y) * w + (v + y) * z.
Apply mul_SNo_Lt (x + y) w (v + y) z Lxy Hw1 Lvy Hz to the current goal.
We will prove v + y < x + y.
Apply add_SNo_Lt1 v y x Hv1 Hy Hx to the current goal.
An exact proof term for the current goal is Hv3.
We will prove z < w.
An exact proof term for the current goal is Hw3.
We will prove v * z + y * z + x * w + y * w (u + v * w) + y * z + y * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) (x * w + y * w) Lvz Lyz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
We will prove (v * z + y * z) + x * w + y * w (u + v * w) + y * z + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw (from left to right).
We will prove (v * z + x * w) + y * z + y * w (u + v * w) + y * z + y * w.
Apply add_SNo_Le1 (v * z + x * w) (y * z + y * w) (u + v * w) (SNo_add_SNo (v * z) (x * w) Lvz Lxw) (SNo_add_SNo (y * z) (y * w) Lyz Lyw) Luvw to the current goal.
We will prove v * z + x * w u + v * w.
An exact proof term for the current goal is Hvw.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoL z.
Assume Hvw: v * z + x * w u + v * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lvy: SNo (v + y).
An exact proof term for the current goal is SNo_add_SNo v y Hv1 Hy.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luyz: SNo (u + y * z).
An exact proof term for the current goal is SNo_add_SNo u (y * z) Hu1 Lyz.
We prove the intermediate claim Lvwyw: SNo (v * w + y * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (y * w) Lvw Lyw.
We will prove (x + y) * z < u + y * z.
Apply add_SNo_Lt1_cancel ((x + y) * z) (v * w + y * w) (u + y * z) Lxyz Lvwyw Luyz to the current goal.
We will prove (x + y) * z + v * w + y * w < (u + y * z) + v * w + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (y * z) (v * w) (y * w) Hu1 Lyz Lvw Lyw (from left to right).
We will prove (x + y) * z + v * w + y * w < (u + v * w) + y * z + y * w.
Apply SNoLtLe_tra ((x + y) * z + v * w + y * w) (v * z + y * z + x * w + y * w) ((u + v * w) + y * z + y * w) (SNo_add_SNo_3 ((x + y) * z) (v * w) (y * w) Lxyz Lvw Lyw) (SNo_add_SNo_4 (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw) (SNo_add_SNo_3 (u + v * w) (y * z) (y * w) Luvw Lyz Lyw) to the current goal.
We will prove (x + y) * z + v * w + y * w < v * z + y * z + x * w + y * w.
rewrite the current goal using IHz w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + v * w + y * w < v * z + y * z + (x + y) * w.
rewrite the current goal using IHxz v (SNoR_SNoS x Hx v Hv) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + (v + y) * w < v * z + y * z + (x + y) * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) ((x + y) * w) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using IHx v (SNoR_SNoS x Hx v Hv) (from right to left).
We will prove (x + y) * z + (v + y) * w < (v + y) * z + (x + y) * w.
Apply mul_SNo_Lt (v + y) z (x + y) w Lvy Hz Lxy Hw1 to the current goal.
We will prove x + y < v + y.
Apply add_SNo_Lt1 x y v Hx Hy Hv1 to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w < z.
An exact proof term for the current goal is Hw3.
We will prove v * z + y * z + x * w + y * w (u + v * w) + y * z + y * w.
rewrite the current goal using add_SNo_assoc (v * z) (y * z) (x * w + y * w) Lvz Lyz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
We will prove (v * z + y * z) + x * w + y * w (u + v * w) + y * z + y * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (y * z) (x * w) (y * w) Lvz Lyz Lxw Lyw (from left to right).
We will prove (v * z + x * w) + y * z + y * w (u + v * w) + y * z + y * w.
Apply add_SNo_Le1 (v * z + x * w) (y * z + y * w) (u + v * w) (SNo_add_SNo (v * z) (x * w) Lvz Lxw) (SNo_add_SNo (y * z) (y * w) Lyz Lyw) Luvw to the current goal.
We will prove v * z + x * w u + v * w.
An exact proof term for the current goal is Hvw.
Assume Hu': u' R2.
Apply ReplE_impred (SNoR (y * z)) (λw ⇒ x * z + w) u' Hu' to the current goal.
Let u be given.
Assume Hu: u SNoR (y * z).
Assume Hu'u: u' = x * z + u.
rewrite the current goal using Hu'u (from left to right).
We will prove (x + y) * z < x * z + u.
Apply SNoR_E (y * z) Lyz u Hu to the current goal.
Assume Hu1 Hu2 Hu3.
rewrite the current goal using add_SNo_com (x * z) u Lxz Hu1 (from left to right).
We will prove (x + y) * z < u + x * z.
Apply mul_SNo_SNoR_interpolate_impred y z Hy Hz u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL y.
Let w be given.
Assume Hw: w SNoR z.
Assume Hvw: v * z + y * w u + v * w.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lxv: SNo (x + v).
An exact proof term for the current goal is SNo_add_SNo x v Hx Hv1.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luxz: SNo (u + x * z).
An exact proof term for the current goal is SNo_add_SNo u (x * z) Hu1 Lxz.
We prove the intermediate claim Lvwxw: SNo (v * w + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (x * w) Lvw Lxw.
We will prove (x + y) * z < u + x * z.
Apply add_SNo_Lt1_cancel ((x + y) * z) (v * w + x * w) (u + x * z) Lxyz Lvwxw Luxz to the current goal.
We will prove (x + y) * z + v * w + x * w < (u + x * z) + v * w + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (x * z) (v * w) (x * w) Hu1 Lxz Lvw Lxw (from left to right).
We will prove (x + y) * z + v * w + x * w < (u + v * w) + x * z + x * w.
Apply SNoLtLe_tra ((x + y) * z + v * w + x * w) (v * z + x * z + x * w + y * w) ((u + v * w) + x * z + x * w) (SNo_add_SNo_3 ((x + y) * z) (v * w) (x * w) Lxyz Lvw Lxw) (SNo_add_SNo_4 (v * z) (x * z) (x * w) (y * w) Lvz Lxz Lxw Lyw) (SNo_add_SNo_3 (u + v * w) (x * z) (x * w) Luvw Lxz Lxw) to the current goal.
We will prove (x + y) * z + v * w + x * w < v * z + x * z + x * w + y * w.
rewrite the current goal using IHz w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + v * w + x * w < v * z + x * z + (x + y) * w.
rewrite the current goal using add_SNo_com (v * w) (x * w) Lvw Lxw (from left to right).
rewrite the current goal using IHyz v (SNoL_SNoS y Hy v Hv) w (SNoR_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + (x + v) * w < v * z + x * z + (x + y) * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) ((x + y) * w) Lvz Lxz Lxyw (from left to right).
We will prove (x + y) * z + (x + v) * w < (v * z + x * z) + (x + y) * w.
rewrite the current goal using add_SNo_com (v * z) (x * z) Lvz Lxz (from left to right).
We will prove (x + y) * z + (x + v) * w < (x * z + v * z) + (x + y) * w.
rewrite the current goal using IHy v (SNoL_SNoS y Hy v Hv) (from right to left).
We will prove (x + y) * z + (x + v) * w < (x + v) * z + (x + y) * w.
rewrite the current goal using add_SNo_com ((x + y) * z) ((x + v) * w) Lxyz (SNo_mul_SNo (x + v) w Lxv Hw1) (from left to right).
rewrite the current goal using add_SNo_com ((x + v) * z) ((x + y) * w) (SNo_mul_SNo (x + v) z Lxv Hz) Lxyw (from left to right).
We will prove (x + v) * w + (x + y) * z < (x + y) * w + (x + v) * z.
Apply mul_SNo_Lt (x + y) w (x + v) z Lxy Hw1 Lxv Hz to the current goal.
We will prove x + v < x + y.
Apply add_SNo_Lt2 x v y Hx Hv1 Hy to the current goal.
An exact proof term for the current goal is Hv3.
We will prove z < w.
An exact proof term for the current goal is Hw3.
We will prove v * z + x * z + x * w + y * w (u + v * w) + x * z + x * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) (x * w + y * w) Lvz Lxz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
We will prove (v * z + x * z) + x * w + y * w (u + v * w) + x * z + x * w.
rewrite the current goal using add_SNo_com (x * w) (y * w) Lxw Lyw (from left to right).
We will prove (v * z + x * z) + y * w + x * w (u + v * w) + x * z + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (x * z) (y * w) (x * w) Lvz Lxz Lyw Lxw (from left to right).
We will prove (v * z + y * w) + x * z + x * w (u + v * w) + x * z + x * w.
Apply add_SNo_Le1 (v * z + y * w) (x * z + x * w) (u + v * w) (SNo_add_SNo (v * z) (y * w) Lvz Lyw) (SNo_add_SNo (x * z) (x * w) Lxz Lxw) Luvw to the current goal.
We will prove v * z + y * w u + v * w.
An exact proof term for the current goal is Hvw.
Let v be given.
Assume Hv: v SNoR y.
Let w be given.
Assume Hw: w SNoL z.
Assume Hvw: v * z + y * w u + v * w.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Hv1 Hw1.
We prove the intermediate claim Lxv: SNo (x + v).
An exact proof term for the current goal is SNo_add_SNo x v Hx Hv1.
We prove the intermediate claim Luvw: SNo (u + v * w).
An exact proof term for the current goal is SNo_add_SNo u (v * w) Hu1 Lvw.
We prove the intermediate claim Lxyw: SNo ((x + y) * w).
An exact proof term for the current goal is SNo_mul_SNo (x + y) w Lxy Hw1.
We prove the intermediate claim Lvz: SNo (v * z).
An exact proof term for the current goal is SNo_mul_SNo v z Hv1 Hz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_mul_SNo x w Hx Hw1.
We prove the intermediate claim Lyw: SNo (y * w).
An exact proof term for the current goal is SNo_mul_SNo y w Hy Hw1.
We prove the intermediate claim Luxz: SNo (u + x * z).
An exact proof term for the current goal is SNo_add_SNo u (x * z) Hu1 Lxz.
We prove the intermediate claim Lvwxw: SNo (v * w + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * w) (x * w) Lvw Lxw.
We prove the intermediate claim Lvzxw: SNo (v * z + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * z) (x * w) Lvz Lxw.
We will prove (x + y) * z < u + x * z.
Apply add_SNo_Lt1_cancel ((x + y) * z) (v * w + x * w) (u + x * z) Lxyz Lvwxw Luxz to the current goal.
We will prove (x + y) * z + v * w + x * w < (u + x * z) + v * w + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid u (x * z) (v * w) (x * w) Hu1 Lxz Lvw Lxw (from left to right).
We will prove (x + y) * z + v * w + x * w < (u + v * w) + x * z + x * w.
Apply SNoLtLe_tra ((x + y) * z + v * w + x * w) (v * z + x * z + x * w + y * w) ((u + v * w) + x * z + x * w) (SNo_add_SNo_3 ((x + y) * z) (v * w) (x * w) Lxyz Lvw Lxw) (SNo_add_SNo_4 (v * z) (x * z) (x * w) (y * w) Lvz Lxz Lxw Lyw) (SNo_add_SNo_3 (u + v * w) (x * z) (x * w) Luvw Lxz Lxw) to the current goal.
We will prove (x + y) * z + v * w + x * w < v * z + x * z + x * w + y * w.
rewrite the current goal using IHz w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + v * w + x * w < v * z + x * z + (x + y) * w.
rewrite the current goal using add_SNo_com (v * w) (x * w) Lvw Lxw (from left to right).
We will prove (x + y) * z + x * w + v * w < v * z + x * z + (x + y) * w.
rewrite the current goal using IHyz v (SNoR_SNoS y Hy v Hv) w (SNoL_SNoS z Hz w Hw) (from right to left).
We will prove (x + y) * z + (x + v) * w < v * z + x * z + (x + y) * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) ((x + y) * w) Lvz Lxz Lxyw (from left to right).
We will prove (x + y) * z + (x + v) * w < (v * z + x * z) + (x + y) * w.
rewrite the current goal using add_SNo_com (v * z) (x * z) Lvz Lxz (from left to right).
rewrite the current goal using IHy v (SNoR_SNoS y Hy v Hv) (from right to left).
We will prove (x + y) * z + (x + v) * w < (x + v) * z + (x + y) * w.
rewrite the current goal using add_SNo_com ((x + v) * z) ((x + y) * w) (SNo_mul_SNo (x + v) z Lxv Hz) Lxyw (from left to right).
rewrite the current goal using add_SNo_com ((x + y) * z) ((x + v) * w) Lxyz (SNo_mul_SNo (x + v) w Lxv Hw1) (from left to right).
We will prove (x + v) * w + (x + y) * z < (x + y) * w + (x + v) * z.
rewrite the current goal using add_SNo_com ((x + v) * w) ((x + y) * z) (SNo_mul_SNo (x + v) w Lxv Hw1) Lxyz (from left to right).
rewrite the current goal using add_SNo_com ((x + y) * w) ((x + v) * z) Lxyw (SNo_mul_SNo (x + v) z Lxv Hz) (from left to right).
We will prove (x + y) * z + (x + v) * w < (x + v) * z + (x + y) * w.
Apply mul_SNo_Lt (x + v) z (x + y) w Lxv Hz Lxy Hw1 to the current goal.
We will prove x + y < x + v.
Apply add_SNo_Lt2 x y v Hx Hy Hv1 to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w < z.
An exact proof term for the current goal is Hw3.
We will prove v * z + x * z + x * w + y * w (u + v * w) + x * z + x * w.
rewrite the current goal using add_SNo_assoc (v * z) (x * z) (x * w + y * w) Lvz Lxz (SNo_add_SNo (x * w) (y * w) Lxw Lyw) (from left to right).
rewrite the current goal using add_SNo_com (y * w) (x * w) Lyw Lxw (from right to left).
We will prove (v * z + x * z) + y * w + x * w (u + v * w) + x * z + x * w.
rewrite the current goal using add_SNo_com_4_inner_mid (v * z) (x * z) (y * w) (x * w) Lvz Lxz Lyw Lxw (from left to right).
We will prove (v * z + y * w) + x * z + x * w (u + v * w) + x * z + x * w.
Apply add_SNo_Le1 (v * z + y * w) (x * z + x * w) (u + v * w) (SNo_add_SNo (v * z) (y * w) Lvz Lyw) (SNo_add_SNo (x * z) (x * w) Lxz Lxw) Luvw to the current goal.
We will prove v * z + y * w u + v * w.
An exact proof term for the current goal is Hvw.
Theorem. (mul_SNo_distrL) The following is provable:
∀x y z, SNo xSNo ySNo zx * (y + z) = x * y + x * z
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
Use transitivity with ((y + z) * x), and (y * x + z * x).
An exact proof term for the current goal is mul_SNo_com x (y + z) Hx (SNo_add_SNo y z Hy Hz).
An exact proof term for the current goal is mul_SNo_distrR y z x Hy Hz Hx.
We will prove y * x + z * x = x * y + x * z.
Use f_equal.
An exact proof term for the current goal is mul_SNo_com y x Hy Hx.
An exact proof term for the current goal is mul_SNo_com z x Hz Hx.
Beginning of Section mul_SNo_assoc_lems
Variable M : setsetset
Hypothesis DL : ∀x y z, SNo xSNo ySNo zx * (y + z) = x * y + x * z
Hypothesis DR : ∀x y z, SNo xSNo ySNo z(x + y) * z = x * z + y * z
Hypothesis IL : ∀x y, SNo xSNo yuSNoL (x * y), ∀p : prop, (vSNoL x, wSNoL y, u + v * w v * y + x * wp)(vSNoR x, wSNoR y, u + v * w v * y + x * wp)p
Hypothesis IR : ∀x y, SNo xSNo yuSNoR (x * y), ∀p : prop, (vSNoL x, wSNoR y, v * y + x * w u + v * wp)(vSNoR x, wSNoL y, v * y + x * w u + v * wp)p
Hypothesis M_Lt : ∀x y u v, SNo xSNo ySNo uSNo vu < xv < yu * y + x * v < x * y + u * v
Hypothesis M_Le : ∀x y u v, SNo xSNo ySNo uSNo vu xv yu * y + x * v x * y + u * v
Theorem. (mul_SNo_assoc_lem1) The following is provable:
∀x y z, SNo xSNo ySNo z(uSNoS_ (SNoLev x), u * (y * z) = (u * y) * z)(vSNoS_ (SNoLev y), x * (v * z) = (x * v) * z)(wSNoS_ (SNoLev z), x * (y * w) = (x * y) * w)(uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), u * (v * z) = (u * v) * z)(uSNoS_ (SNoLev x), wSNoS_ (SNoLev z), u * (y * w) = (u * y) * w)(vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), x * (v * w) = (x * v) * w)(uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), u * (v * w) = (u * v) * w)∀L, (uL, ∀q : prop, (vSNoL x, wSNoL (y * z), u = v * (y * z) + x * w + - v * wq)(vSNoR x, wSNoR (y * z), u = v * (y * z) + x * w + - v * wq)q)uL, u < (x * y) * z
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
Assume IHx IHy IHz IHxy IHxz IHyz IHxyz.
Let L be given.
Assume HLE.
Let u be given.
Assume Hu: u L.
We will prove u < (x * y) * z.
We prove the intermediate claim Lxy: SNo (x * y).
An exact proof term for the current goal is SNo_M x y Hx Hy.
We prove the intermediate claim Lyz: SNo (y * z).
An exact proof term for the current goal is SNo_M y z Hy Hz.
We prove the intermediate claim Lxyz2: SNo ((x * y) * z).
An exact proof term for the current goal is SNo_M (x * y) z Lxy Hz.
We prove the intermediate claim L1: vSNoS_ (SNoLev x), ∀w, SNo ww1SNoS_ (SNoLev y), w2SNoS_ (SNoLev z), u = v * (y * z) + x * w + - v * wv * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2)(v * y + x * w1) * z + (x * y + v * w1) * w2 < (x * y + v * w1) * z + (v * y + x * w1) * w2u < (x * y) * z.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
Let w1 be given.
Assume Hw1.
Let w2 be given.
Assume Hw2.
Assume Hue H1 H2.
Apply SNoS_E2 (SNoLev x) (SNoLev_ordinal x Hx) v Hv to the current goal.
Assume Hv1 Hv2 Hv3 Hv4.
Apply SNoS_E2 (SNoLev y) (SNoLev_ordinal y Hy) w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 Hw14.
Apply SNoS_E2 (SNoLev z) (SNoLev_ordinal z Hz) w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 Hw24.
We prove the intermediate claim Lvyz: SNo (v * (y * z)).
An exact proof term for the current goal is SNo_M v (y * z) Hv3 Lyz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_M x w Hx Hw.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw13.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_M v w Hv3 Hw.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv3 Hw13.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv3 Hy.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw13 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw23.
We prove the intermediate claim Lvw1z: SNo (v * (w1 * z)).
An exact proof term for the current goal is SNo_M v (w1 * z) Hv3 Lw1z.
We prove the intermediate claim Lvyw2: SNo (v * (y * w2)).
An exact proof term for the current goal is SNo_M v (y * w2) Hv3 Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw13 Hw23.
We prove the intermediate claim Lxw1w2: SNo (x * (w1 * w2)).
An exact proof term for the current goal is SNo_M x (w1 * w2) Hx Lw1w2.
We prove the intermediate claim Lxw1z: SNo (x * (w1 * z)).
An exact proof term for the current goal is SNo_M x (w1 * z) Hx Lw1z.
We prove the intermediate claim Lxyw2: SNo (x * (y * w2)).
An exact proof term for the current goal is SNo_M x (y * w2) Hx Lyw2.
We prove the intermediate claim Lvyzxw: SNo (v * (y * z) + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * (y * z)) (x * w) Lvyz Lxw.
We prove the intermediate claim Lxyzvw: SNo ((x * y) * z + v * w).
An exact proof term for the current goal is SNo_add_SNo ((x * y) * z) (v * w) Lxyz2 Lvw.
We prove the intermediate claim Lvw1w2: SNo (v * (w1 * w2)).
An exact proof term for the current goal is SNo_M v (w1 * w2) Hv3 Lw1w2.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Hw Lw1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv3 Lww1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv3 (SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2).
We prove the intermediate claim Lvwvw1w2: SNo (v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo (v * w) (v * (w1 * w2)) Lvw Lvw1w2.
We prove the intermediate claim Lvyzxw1zxyw2vwvw1w2: SNo (v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo_4 (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) (v * w + v * (w1 * w2)) Lvyz Lxw1z Lxyw2 Lvwvw1w2.
We prove the intermediate claim Lvw1zvyw2xw1w2: SNo (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo_3 (v * (w1 * z)) (v * (y * w2)) (x * (w1 * w2)) Lvw1z Lvyw2 Lxw1w2.
We will prove u < (x * y) * z.
rewrite the current goal using Hue (from left to right).
We will prove (v * (y * z)) + (x * w) + - (v * w) < (x * y) * z.
Apply add_SNo_minus_Lt1b3 (v * (y * z)) (x * w) (v * w) ((x * y) * z) Lvyz Lxw Lvw Lxyz2 to the current goal.
We will prove v * (y * z) + x * w < (x * y) * z + v * w.
Apply add_SNo_Lt1_cancel (v * (y * z) + x * w) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) ((x * y) * z + v * w) Lvyzxw Lvw1zvyw2xw1w2 Lxyzvw to the current goal.
We will prove (v * (y * z) + x * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) < ((x * y) * z + v * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
Apply SNoLeLt_tra ((v * (y * z) + x * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) (v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2)) (((x * y) * z + v * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) to the current goal.
Apply SNo_add_SNo to the current goal.
An exact proof term for the current goal is Lvyzxw.
An exact proof term for the current goal is Lvw1zvyw2xw1w2.
An exact proof term for the current goal is Lvyzxw1zxyw2vwvw1w2.
Apply SNo_add_SNo to the current goal.
An exact proof term for the current goal is Lxyzvw.
An exact proof term for the current goal is Lvw1zvyw2xw1w2.
We will prove (v * (y * z) + x * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
rewrite the current goal using add_SNo_assoc (v * (y * z)) (x * w) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lvyz Lxw Lvw1zvyw2xw1w2 (from right to left).
We will prove v * (y * z) + x * w + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
Apply add_SNo_Le2 (v * (y * z)) (x * w + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) (x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2)) Lvyz (SNo_add_SNo (x * w) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lxw Lvw1zvyw2xw1w2) (SNo_add_SNo_4 (x * (w1 * z)) (x * (y * w2)) (v * w) (v * (w1 * w2)) Lxw1z Lxyw2 Lvw Lvw1w2) to the current goal.
We will prove x * w + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
rewrite the current goal using add_SNo_assoc (v * (w1 * z)) (v * (y * w2)) (x * (w1 * w2)) Lvw1z Lvyw2 Lxw1w2 (from left to right).
rewrite the current goal using DL v (w1 * z) (y * w2) Hv3 Lw1z Lyw2 (from right to left).
We will prove x * w + v * (w1 * z + y * w2) + x * (w1 * w2) x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
rewrite the current goal using add_SNo_com_3_0_1 (x * w) (v * (w1 * z + y * w2)) (x * (w1 * w2)) Lxw Lvw1zyw2 Lxw1w2 (from left to right).
We will prove v * (w1 * z + y * w2) + x * w + x * (w1 * w2) x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
rewrite the current goal using DL x w (w1 * w2) Hx Hw Lw1w2 (from right to left).
We will prove v * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
rewrite the current goal using DL v w (w1 * w2) Hv3 Hw Lw1w2 (from right to left).
We will prove v * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z) + x * (y * w2) + v * (w + w1 * w2).
rewrite the current goal using add_SNo_assoc (x * (w1 * z)) (x * (y * w2)) (v * (w + w1 * w2)) Lxw1z Lxyw2 Lvww1w2 (from left to right).
rewrite the current goal using DL x (w1 * z) (y * w2) Hx Lw1z Lyw2 (from right to left).
We will prove v * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2).
An exact proof term for the current goal is H1.
We will prove v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) < ((x * y) * z + v * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
rewrite the current goal using add_SNo_com ((x * y) * z) (v * w) Lxyz2 Lvw (from left to right).
rewrite the current goal using add_SNo_assoc (v * w) ((x * y) * z) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lvw Lxyz2 Lvw1zvyw2xw1w2 (from right to left).
We will prove v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) < v * w + (x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
rewrite the current goal using add_SNo_rotate_5_2 (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) (v * w) (v * (w1 * w2)) Lvyz Lxw1z Lxyw2 Lvw Lvw1w2 (from left to right).
We will prove v * w + v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2) < v * w + (x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
Apply add_SNo_Lt2 (v * w) (v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2)) ((x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lvw (SNo_add_SNo_4 (v * (w1 * w2)) (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) Lvw1w2 Lvyz Lxw1z Lxyw2) (SNo_add_SNo_4 ((x * y) * z) (v * (w1 * z)) (v * (y * w2)) (x * (w1 * w2)) Lxyz2 Lvw1z Lvyw2 Lxw1w2) to the current goal.
We will prove v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2) < (x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
rewrite the current goal using add_SNo_assoc ((x * y) * z) (v * (w1 * z)) (v * (y * w2) + x * (w1 * w2)) Lxyz2 Lvw1z (SNo_add_SNo (v * (y * w2)) (x * (w1 * w2)) Lvyw2 Lxw1w2) (from left to right).
We will prove v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2) < ((x * y) * z + v * (w1 * z)) + v * (y * w2) + x * (w1 * w2).
rewrite the current goal using IHxz v Hv w2 Hw2 (from left to right).
rewrite the current goal using IHyz w1 Hw1 w2 Hw2 (from left to right).
We will prove v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2) < ((x * y) * z + v * (w1 * z)) + (v * y) * w2 + (x * w1) * w2.
rewrite the current goal using DR (v * y) (x * w1) w2 Lvy Lxw1 Hw23 (from right to left).
We will prove v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2) < ((x * y) * z + v * (w1 * z)) + (v * y + x * w1) * w2.
rewrite the current goal using IHxy v Hv w1 Hw1 (from left to right).
rewrite the current goal using DR (x * y) (v * w1) z Lxy Lvw1 Hz (from right to left).
We will prove v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2) < (x * y + v * w1) * z + (v * y + x * w1) * w2.
rewrite the current goal using add_SNo_rotate_4_1 (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) (v * (w1 * w2)) Lvyz Lxw1z Lxyw2 Lvw1w2 (from right to left).
We will prove v * (y * z) + x * (w1 * z) + x * (y * w2) + v * (w1 * w2) < (x * y + v * w1) * z + (v * y + x * w1) * w2.
rewrite the current goal using add_SNo_assoc (v * (y * z)) (x * (w1 * z)) (x * (y * w2) + v * (w1 * w2)) Lvyz Lxw1z (SNo_add_SNo (x * (y * w2)) (v * (w1 * w2)) Lxyw2 Lvw1w2) (from left to right).
We will prove (v * (y * z) + x * (w1 * z)) + x * (y * w2) + v * (w1 * w2) < (x * y + v * w1) * z + (v * y + x * w1) * w2.
rewrite the current goal using IHx v Hv (from left to right).
rewrite the current goal using IHy w1 Hw1 (from left to right).
rewrite the current goal using IHz w2 Hw2 (from left to right).
rewrite the current goal using IHxyz v Hv w1 Hw1 w2 Hw2 (from left to right).
We will prove ((v * y) * z + (x * w1) * z) + (x * y) * w2 + (v * w1) * w2 < (x * y + v * w1) * z + (v * y + x * w1) * w2.
rewrite the current goal using DR (v * y) (x * w1) z Lvy Lxw1 Hz (from right to left).
rewrite the current goal using DR (x * y) (v * w1) w2 Lxy Lvw1 Hw23 (from right to left).
We will prove (v * y + x * w1) * z + (x * y + v * w1) * w2 < (x * y + v * w1) * z + (v * y + x * w1) * w2.
An exact proof term for the current goal is H2.
Apply HLE u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoL (y * z).
Assume Hue: u = v * (y * z) + x * w + - v * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
We prove the intermediate claim Lw: SNo w.
Apply SNoL_E (y * z) Lyz w Hw to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
Apply IL y z Hy Hz w Hw to the current goal.
Let w1 be given.
Assume Hw1: w1 SNoL y.
Let w2 be given.
Assume Hw2: w2 SNoL z.
Assume Hwl: w + w1 * w2 w1 * z + y * w2.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
Apply L1 v (SNoL_SNoS x Hx v Hv) w Lw w1 (SNoL_SNoS y Hy w1 Hw1) w2 (SNoL_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove v * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2).
Apply M_Le x (w1 * z + y * w2) v (w + w1 * w2) Hx Lw1zyw2 Hv1 Lww1w2 to the current goal.
We will prove v x.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w + w1 * w2 w1 * z + y * w2.
An exact proof term for the current goal is Hwl.
We will prove (v * y + x * w1) * z + (x * y + v * w1) * w2 < (x * y + v * w1) * z + (v * y + x * w1) * w2.
Apply M_Lt (x * y + v * w1) z (v * y + x * w1) w2 (SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11)) Hz (SNo_add_SNo (v * y) (x * w1) (SNo_M v y Hv1 Hy) (SNo_M x w1 Hx Hw11)) Hw21 to the current goal.
We will prove v * y + x * w1 < x * y + v * w1.
Apply M_Lt x y v w1 Hx Hy Hv1 Hw11 to the current goal.
We will prove v < x.
An exact proof term for the current goal is Hv3.
We will prove w1 < y.
An exact proof term for the current goal is Hw13.
We will prove w2 < z.
An exact proof term for the current goal is Hw23.
Let w1 be given.
Assume Hw1: w1 SNoR y.
Let w2 be given.
Assume Hw2: w2 SNoR z.
Assume Hwl: w + w1 * w2 w1 * z + y * w2.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv1 Hy.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw11.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
Apply L1 v (SNoL_SNoS x Hx v Hv) w Lw w1 (SNoR_SNoS y Hy w1 Hw1) w2 (SNoR_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove v * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2).
Apply M_Le x (w1 * z + y * w2) v (w + w1 * w2) Hx Lw1zyw2 Hv1 Lww1w2 to the current goal.
We will prove v x.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w + w1 * w2 w1 * z + y * w2.
An exact proof term for the current goal is Hwl.
We will prove (v * y + x * w1) * z + (x * y + v * w1) * w2 < (x * y + v * w1) * z + (v * y + x * w1) * w2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) (SNo_M v y Hv1 Hy) (SNo_M x w1 Hx Hw11).
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11).
We prove the intermediate claim Lvyxw1z: SNo ((v * y + x * w1) * z).
An exact proof term for the current goal is SNo_M (v * y + x * w1) z Lvyxw1 Hz.
We prove the intermediate claim Lxyvw1z: SNo ((x * y + v * w1) * z).
An exact proof term for the current goal is SNo_M (x * y + v * w1) z Lxyvw1 Hz.
We prove the intermediate claim Lvyxw1w2: SNo ((v * y + x * w1) * w2).
An exact proof term for the current goal is SNo_M (v * y + x * w1) w2 Lvyxw1 Hw21.
We prove the intermediate claim Lxyvw1w2: SNo ((x * y + v * w1) * w2).
An exact proof term for the current goal is SNo_M (x * y + v * w1) w2 Lxyvw1 Hw21.
rewrite the current goal using add_SNo_com ((v * y + x * w1) * z) ((x * y + v * w1) * w2) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using add_SNo_com ((x * y + v * w1) * z) ((v * y + x * w1) * w2) Lxyvw1z Lvyxw1w2 (from left to right).
We will prove (x * y + v * w1) * w2 + (v * y + x * w1) * z < (v * y + x * w1) * w2 + (x * y + v * w1) * z.
Apply M_Lt (v * y + x * w1) w2 (x * y + v * w1) z Lvyxw1 Hw21 Lxyvw1 Hz to the current goal.
We will prove x * y + v * w1 < v * y + x * w1.
rewrite the current goal using add_SNo_com (x * y) (v * w1) Lxy Lvw1 (from left to right).
rewrite the current goal using add_SNo_com (v * y) (x * w1) Lvy Lxw1 (from left to right).
Apply M_Lt x w1 v y Hx Hw11 Hv1 Hy to the current goal.
We will prove v < x.
An exact proof term for the current goal is Hv3.
We will prove y < w1.
An exact proof term for the current goal is Hw13.
We will prove z < w2.
An exact proof term for the current goal is Hw23.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoR (y * z).
Assume Hue: u = v * (y * z) + x * w + - v * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
We prove the intermediate claim Lw: SNo w.
Apply SNoR_E (y * z) Lyz w Hw to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_M v w Hv1 Lw.
Apply IR y z Hy Hz w Hw to the current goal.
Let w1 be given.
Assume Hw1: w1 SNoL y.
Let w2 be given.
Assume Hw2: w2 SNoR z.
Assume Hwl: w1 * z + y * w2 w + w1 * w2.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv1 Hy.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw11.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
We prove the intermediate claim Lxww1w2: SNo (x * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M x (w + w1 * w2) Hx Lww1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv1 Lww1w2.
We prove the intermediate claim Lvw1w2: SNo (v * (w1 * w2)).
An exact proof term for the current goal is SNo_M v (w1 * w2) Hv1 Lw1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv1 Lw1zyw2.
We prove the intermediate claim Lvwvw1w2: SNo (v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo (v * w) (v * (w1 * w2)) Lvw Lvw1w2.
We prove the intermediate claim Lxw1zyw2: SNo (x * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M x (w1 * z + y * w2) Hx Lw1zyw2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) Lvy Lxw1.
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy Lvw1.
We prove the intermediate claim Lvyxw1z: SNo ((v * y + x * w1) * z).
An exact proof term for the current goal is SNo_M (v * y + x * w1) z Lvyxw1 Hz.
We prove the intermediate claim Lxyvw1z: SNo ((x * y + v * w1) * z).
An exact proof term for the current goal is SNo_M (x * y + v * w1) z Lxyvw1 Hz.
We prove the intermediate claim Lvyxw1w2: SNo ((v * y + x * w1) * w2).
An exact proof term for the current goal is SNo_M (v * y + x * w1) w2 Lvyxw1 Hw21.
We prove the intermediate claim Lxyvw1w2: SNo ((x * y + v * w1) * w2).
An exact proof term for the current goal is SNo_M (x * y + v * w1) w2 Lxyvw1 Hw21.
Apply L1 v (SNoR_SNoS x Hx v Hv) w Lw w1 (SNoL_SNoS y Hy w1 Hw1) w2 (SNoR_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove v * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2).
rewrite the current goal using add_SNo_com (v * (w1 * z + y * w2)) (x * (w + w1 * w2)) Lvw1zyw2 Lxww1w2 (from left to right).
We will prove x * (w + w1 * w2) + v * (w1 * z + y * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2).
rewrite the current goal using add_SNo_com (x * (w1 * z + y * w2)) (v * (w + w1 * w2)) Lxw1zyw2 Lvww1w2 (from left to right).
We will prove x * (w + w1 * w2) + v * (w1 * z + y * w2) v * (w + w1 * w2) + x * (w1 * z + y * w2).
Apply M_Le v (w + w1 * w2) x (w1 * z + y * w2) Hv1 Lww1w2 Hx Lw1zyw2 to the current goal.
We will prove x v.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w1 * z + y * w2 w + w1 * w2.
An exact proof term for the current goal is Hwl.
We will prove (v * y + x * w1) * z + (x * y + v * w1) * w2 < (x * y + v * w1) * z + (v * y + x * w1) * w2.
rewrite the current goal using add_SNo_com ((v * y + x * w1) * z) ((x * y + v * w1) * w2) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using add_SNo_com ((x * y + v * w1) * z) ((v * y + x * w1) * w2) Lxyvw1z Lvyxw1w2 (from left to right).
We will prove (x * y + v * w1) * w2 + (v * y + x * w1) * z < (v * y + x * w1) * w2 + (x * y + v * w1) * z.
Apply M_Lt (v * y + x * w1) w2 (x * y + v * w1) z Lvyxw1 Hw21 Lxyvw1 Hz to the current goal.
We will prove x * y + v * w1 < v * y + x * w1.
Apply M_Lt v y x w1 Hv1 Hy Hx Hw11 to the current goal.
We will prove x < v.
An exact proof term for the current goal is Hv3.
We will prove w1 < y.
An exact proof term for the current goal is Hw13.
We will prove z < w2.
An exact proof term for the current goal is Hw23.
Let w1 be given.
Assume Hw1: w1 SNoR y.
Let w2 be given.
Assume Hw2: w2 SNoL z.
Assume Hwl: w1 * z + y * w2 w + w1 * w2.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv1 Hy.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw11.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
We prove the intermediate claim Lxww1w2: SNo (x * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M x (w + w1 * w2) Hx Lww1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv1 Lww1w2.
We prove the intermediate claim Lvw1w2: SNo (v * (w1 * w2)).
An exact proof term for the current goal is SNo_M v (w1 * w2) Hv1 Lw1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv1 Lw1zyw2.
We prove the intermediate claim Lvwvw1w2: SNo (v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo (v * w) (v * (w1 * w2)) Lvw Lvw1w2.
We prove the intermediate claim Lxw1zyw2: SNo (x * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M x (w1 * z + y * w2) Hx Lw1zyw2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) Lvy Lxw1.
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy Lvw1.
We prove the intermediate claim Lvyxw1z: SNo ((v * y + x * w1) * z).
An exact proof term for the current goal is SNo_M (v * y + x * w1) z Lvyxw1 Hz.
We prove the intermediate claim Lxyvw1z: SNo ((x * y + v * w1) * z).
An exact proof term for the current goal is SNo_M (x * y + v * w1) z Lxyvw1 Hz.
We prove the intermediate claim Lvyxw1w2: SNo ((v * y + x * w1) * w2).
An exact proof term for the current goal is SNo_M (v * y + x * w1) w2 Lvyxw1 Hw21.
We prove the intermediate claim Lxyvw1w2: SNo ((x * y + v * w1) * w2).
An exact proof term for the current goal is SNo_M (x * y + v * w1) w2 Lxyvw1 Hw21.
Apply L1 v (SNoR_SNoS x Hx v Hv) w Lw w1 (SNoR_SNoS y Hy w1 Hw1) w2 (SNoL_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove v * (w1 * z + y * w2) + x * (w + w1 * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2).
rewrite the current goal using add_SNo_com (v * (w1 * z + y * w2)) (x * (w + w1 * w2)) Lvw1zyw2 Lxww1w2 (from left to right).
We will prove x * (w + w1 * w2) + v * (w1 * z + y * w2) x * (w1 * z + y * w2) + v * (w + w1 * w2).
rewrite the current goal using add_SNo_com (x * (w1 * z + y * w2)) (v * (w + w1 * w2)) Lxw1zyw2 Lvww1w2 (from left to right).
We will prove x * (w + w1 * w2) + v * (w1 * z + y * w2) v * (w + w1 * w2) + x * (w1 * z + y * w2).
Apply M_Le v (w + w1 * w2) x (w1 * z + y * w2) Hv1 Lww1w2 Hx Lw1zyw2 to the current goal.
We will prove x v.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
We will prove w1 * z + y * w2 w + w1 * w2.
An exact proof term for the current goal is Hwl.
We will prove (v * y + x * w1) * z + (x * y + v * w1) * w2 < (x * y + v * w1) * z + (v * y + x * w1) * w2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) (SNo_M v y Hv1 Hy) (SNo_M x w1 Hx Hw11).
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11).
Apply M_Lt (x * y + v * w1) z (v * y + x * w1) w2 Lxyvw1 Hz Lvyxw1 Hw21 to the current goal.
We will prove v * y + x * w1 < x * y + v * w1.
rewrite the current goal using add_SNo_com (x * y) (v * w1) Lxy Lvw1 (from left to right).
rewrite the current goal using add_SNo_com (v * y) (x * w1) Lvy Lxw1 (from left to right).
We will prove x * w1 + v * y < v * w1 + x * y.
Apply M_Lt v w1 x y Hv1 Hw11 Hx Hy to the current goal.
We will prove x < v.
An exact proof term for the current goal is Hv3.
We will prove y < w1.
An exact proof term for the current goal is Hw13.
We will prove w2 < z.
An exact proof term for the current goal is Hw23.
Theorem. (mul_SNo_assoc_lem2) The following is provable:
∀x y z, SNo xSNo ySNo z(uSNoS_ (SNoLev x), u * (y * z) = (u * y) * z)(vSNoS_ (SNoLev y), x * (v * z) = (x * v) * z)(wSNoS_ (SNoLev z), x * (y * w) = (x * y) * w)(uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), u * (v * z) = (u * v) * z)(uSNoS_ (SNoLev x), wSNoS_ (SNoLev z), u * (y * w) = (u * y) * w)(vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), x * (v * w) = (x * v) * w)(uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), u * (v * w) = (u * v) * w)∀R, (uR, ∀q : prop, (vSNoL x, wSNoR (y * z), u = v * (y * z) + x * w + - v * wq)(vSNoR x, wSNoL (y * z), u = v * (y * z) + x * w + - v * wq)q)uR, (x * y) * z < u
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
Assume IHx IHy IHz IHxy IHxz IHyz IHxyz.
Let R be given.
Assume HRE.
Let u be given.
Assume Hu: u R.
We will prove (x * y) * z < u.
We prove the intermediate claim Lxy: SNo (x * y).
An exact proof term for the current goal is SNo_M x y Hx Hy.
We prove the intermediate claim Lyz: SNo (y * z).
An exact proof term for the current goal is SNo_M y z Hy Hz.
We prove the intermediate claim Lxyz2: SNo ((x * y) * z).
An exact proof term for the current goal is SNo_M (x * y) z Lxy Hz.
We prove the intermediate claim L2: vSNoS_ (SNoLev x), ∀w, SNo ww1SNoS_ (SNoLev y), w2SNoS_ (SNoLev z), u = v * (y * z) + x * w + - v * wx * (w1 * z + y * w2) + v * (w + w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2)(x * y + v * w1) * z + (v * y + x * w1) * w2 < (v * y + x * w1) * z + (x * y + v * w1) * w2(x * y) * z < u.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
Let w1 be given.
Assume Hw1.
Let w2 be given.
Assume Hw2.
Assume Hue H1 H2.
Apply SNoS_E2 (SNoLev x) (SNoLev_ordinal x Hx) v Hv to the current goal.
Assume Hv1 Hv2 Hv3 Hv4.
Apply SNoS_E2 (SNoLev y) (SNoLev_ordinal y Hy) w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 Hw14.
Apply SNoS_E2 (SNoLev z) (SNoLev_ordinal z Hz) w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 Hw24.
We prove the intermediate claim Lvyz: SNo (v * (y * z)).
An exact proof term for the current goal is SNo_M v (y * z) Hv3 Lyz.
We prove the intermediate claim Lxw: SNo (x * w).
An exact proof term for the current goal is SNo_M x w Hx Hw.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw13.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_M v w Hv3 Hw.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv3 Hw13.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv3 Hy.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw13 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw23.
We prove the intermediate claim Lvw1z: SNo (v * (w1 * z)).
An exact proof term for the current goal is SNo_M v (w1 * z) Hv3 Lw1z.
We prove the intermediate claim Lvyw2: SNo (v * (y * w2)).
An exact proof term for the current goal is SNo_M v (y * w2) Hv3 Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw13 Hw23.
We prove the intermediate claim Lxw1w2: SNo (x * (w1 * w2)).
An exact proof term for the current goal is SNo_M x (w1 * w2) Hx Lw1w2.
We prove the intermediate claim Lxw1z: SNo (x * (w1 * z)).
An exact proof term for the current goal is SNo_M x (w1 * z) Hx Lw1z.
We prove the intermediate claim Lxyw2: SNo (x * (y * w2)).
An exact proof term for the current goal is SNo_M x (y * w2) Hx Lyw2.
We prove the intermediate claim Lvyzxw: SNo (v * (y * z) + x * w).
An exact proof term for the current goal is SNo_add_SNo (v * (y * z)) (x * w) Lvyz Lxw.
We prove the intermediate claim Lxyzvw: SNo ((x * y) * z + v * w).
An exact proof term for the current goal is SNo_add_SNo ((x * y) * z) (v * w) Lxyz2 Lvw.
We prove the intermediate claim Lvw1w2: SNo (v * (w1 * w2)).
An exact proof term for the current goal is SNo_M v (w1 * w2) Hv3 Lw1w2.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Hw Lw1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv3 Lww1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv3 (SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2).
We prove the intermediate claim Lvwvw1w2: SNo (v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo (v * w) (v * (w1 * w2)) Lvw Lvw1w2.
We prove the intermediate claim Lvyzxw1zxyw2vwvw1w2: SNo (v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo_4 (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) (v * w + v * (w1 * w2)) Lvyz Lxw1z Lxyw2 Lvwvw1w2.
We prove the intermediate claim Lvw1zvyw2xw1w2: SNo (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo_3 (v * (w1 * z)) (v * (y * w2)) (x * (w1 * w2)) Lvw1z Lvyw2 Lxw1w2.
rewrite the current goal using Hue (from left to right).
Apply add_SNo_minus_Lt2b3 (v * (y * z)) (x * w) (v * w) ((x * y) * z) Lvyz Lxw Lvw Lxyz2 to the current goal.
We will prove (x * y) * z + v * w < v * (y * z) + x * w.
Apply add_SNo_Lt1_cancel ((x * y) * z + v * w) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) (v * (y * z) + x * w) Lxyzvw Lvw1zvyw2xw1w2 Lvyzxw to the current goal.
We will prove ((x * y) * z + v * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) < (v * (y * z) + x * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
Apply SNoLtLe_tra (((x * y) * z + v * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) (v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2)) ((v * (y * z) + x * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) to the current goal.
Apply SNo_add_SNo to the current goal.
An exact proof term for the current goal is Lxyzvw.
An exact proof term for the current goal is Lvw1zvyw2xw1w2.
An exact proof term for the current goal is Lvyzxw1zxyw2vwvw1w2.
Apply SNo_add_SNo to the current goal.
An exact proof term for the current goal is Lvyzxw.
An exact proof term for the current goal is Lvw1zvyw2xw1w2.
We will prove ((x * y) * z + v * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) < v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
rewrite the current goal using add_SNo_com ((x * y) * z) (v * w) Lxyz2 Lvw (from left to right).
rewrite the current goal using add_SNo_assoc (v * w) ((x * y) * z) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lvw Lxyz2 Lvw1zvyw2xw1w2 (from right to left).
We will prove v * w + (x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) < v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2).
rewrite the current goal using add_SNo_rotate_5_2 (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) (v * w) (v * (w1 * w2)) Lvyz Lxw1z Lxyw2 Lvw Lvw1w2 (from left to right).
We will prove v * w + (x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) < v * w + v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2).
Apply add_SNo_Lt2 (v * w) ((x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) (v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2)) Lvw (SNo_add_SNo_4 ((x * y) * z) (v * (w1 * z)) (v * (y * w2)) (x * (w1 * w2)) Lxyz2 Lvw1z Lvyw2 Lxw1w2) (SNo_add_SNo_4 (v * (w1 * w2)) (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) Lvw1w2 Lvyz Lxw1z Lxyw2) to the current goal.
We will prove (x * y) * z + v * (w1 * z) + v * (y * w2) + x * (w1 * w2) < v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2).
rewrite the current goal using add_SNo_assoc ((x * y) * z) (v * (w1 * z)) (v * (y * w2) + x * (w1 * w2)) Lxyz2 Lvw1z (SNo_add_SNo (v * (y * w2)) (x * (w1 * w2)) Lvyw2 Lxw1w2) (from left to right).
We will prove ((x * y) * z + v * (w1 * z)) + v * (y * w2) + x * (w1 * w2) < v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2).
rewrite the current goal using IHxz v Hv w2 Hw2 (from left to right).
rewrite the current goal using IHyz w1 Hw1 w2 Hw2 (from left to right).
We will prove ((x * y) * z + v * (w1 * z)) + (v * y) * w2 + (x * w1) * w2 < v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2).
rewrite the current goal using DR (v * y) (x * w1) w2 Lvy Lxw1 Hw23 (from right to left).
We will prove ((x * y) * z + v * (w1 * z)) + (v * y + x * w1) * w2 < v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2).
rewrite the current goal using IHxy v Hv w1 Hw1 (from left to right).
rewrite the current goal using DR (x * y) (v * w1) z Lxy Lvw1 Hz (from right to left).
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < v * (w1 * w2) + v * (y * z) + x * (w1 * z) + x * (y * w2).
rewrite the current goal using add_SNo_rotate_4_1 (v * (y * z)) (x * (w1 * z)) (x * (y * w2)) (v * (w1 * w2)) Lvyz Lxw1z Lxyw2 Lvw1w2 (from right to left).
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < v * (y * z) + x * (w1 * z) + x * (y * w2) + v * (w1 * w2).
rewrite the current goal using add_SNo_assoc (v * (y * z)) (x * (w1 * z)) (x * (y * w2) + v * (w1 * w2)) Lvyz Lxw1z (SNo_add_SNo (x * (y * w2)) (v * (w1 * w2)) Lxyw2 Lvw1w2) (from left to right).
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < (v * (y * z) + x * (w1 * z)) + x * (y * w2) + v * (w1 * w2).
rewrite the current goal using IHx v Hv (from left to right).
rewrite the current goal using IHy w1 Hw1 (from left to right).
rewrite the current goal using IHz w2 Hw2 (from left to right).
rewrite the current goal using IHxyz v Hv w1 Hw1 w2 Hw2 (from left to right).
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < ((v * y) * z + (x * w1) * z) + (x * y) * w2 + (v * w1) * w2.
rewrite the current goal using DR (v * y) (x * w1) z Lvy Lxw1 Hz (from right to left).
rewrite the current goal using DR (x * y) (v * w1) w2 Lxy Lvw1 Hw23 (from right to left).
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < (v * y + x * w1) * z + (x * y + v * w1) * w2.
An exact proof term for the current goal is H2.
We will prove v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) (v * (y * z) + x * w) + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
rewrite the current goal using add_SNo_assoc (v * (y * z)) (x * w) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lvyz Lxw Lvw1zvyw2xw1w2 (from right to left).
We will prove v * (y * z) + x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) v * (y * z) + x * w + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
Apply add_SNo_Le2 (v * (y * z)) (x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2)) (x * w + v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lvyz (SNo_add_SNo_4 (x * (w1 * z)) (x * (y * w2)) (v * w) (v * (w1 * w2)) Lxw1z Lxyw2 Lvw Lvw1w2) (SNo_add_SNo (x * w) (v * (w1 * z) + v * (y * w2) + x * (w1 * w2)) Lxw Lvw1zvyw2xw1w2) to the current goal.
We will prove x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) x * w + v * (w1 * z) + v * (y * w2) + x * (w1 * w2).
rewrite the current goal using add_SNo_assoc (v * (w1 * z)) (v * (y * w2)) (x * (w1 * w2)) Lvw1z Lvyw2 Lxw1w2 (from left to right).
rewrite the current goal using DL v (w1 * z) (y * w2) Hv3 Lw1z Lyw2 (from right to left).
We will prove x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) x * w + v * (w1 * z + y * w2) + x * (w1 * w2).
rewrite the current goal using add_SNo_com_3_0_1 (x * w) (v * (w1 * z + y * w2)) (x * (w1 * w2)) Lxw Lvw1zyw2 Lxw1w2 (from left to right).
We will prove x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) v * (w1 * z + y * w2) + x * w + x * (w1 * w2).
rewrite the current goal using DL x w (w1 * w2) Hx Hw Lw1w2 (from right to left).
We will prove x * (w1 * z) + x * (y * w2) + v * w + v * (w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2).
rewrite the current goal using DL v w (w1 * w2) Hv3 Hw Lw1w2 (from right to left).
We will prove x * (w1 * z) + x * (y * w2) + v * (w + w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2).
rewrite the current goal using add_SNo_assoc (x * (w1 * z)) (x * (y * w2)) (v * (w + w1 * w2)) Lxw1z Lxyw2 Lvww1w2 (from left to right).
rewrite the current goal using DL x (w1 * z) (y * w2) Hx Lw1z Lyw2 (from right to left).
We will prove x * (w1 * z + y * w2) + v * (w + w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2).
An exact proof term for the current goal is H1.
Apply HRE u Hu to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoR (y * z).
Assume Hue: u = v * (y * z) + x * w + - v * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
We prove the intermediate claim Lw: SNo w.
Apply SNoR_E (y * z) Lyz w Hw to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_M v w Hv1 Lw.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv1 Hy.
Apply IR y z Hy Hz w Hw to the current goal.
Let w1 be given.
Assume Hw1: w1 SNoL y.
Let w2 be given.
Assume Hw2: w2 SNoR z.
Assume Hwl: w1 * z + y * w2 w + w1 * w2.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw11.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
We prove the intermediate claim Lxww1w2: SNo (x * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M x (w + w1 * w2) Hx Lww1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv1 Lww1w2.
We prove the intermediate claim Lvw1w2: SNo (v * (w1 * w2)).
An exact proof term for the current goal is SNo_M v (w1 * w2) Hv1 Lw1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv1 Lw1zyw2.
We prove the intermediate claim Lvwvw1w2: SNo (v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo (v * w) (v * (w1 * w2)) Lvw Lvw1w2.
We prove the intermediate claim Lxw1zyw2: SNo (x * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M x (w1 * z + y * w2) Hx Lw1zyw2.
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11).
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) Lvy Lxw1.
We prove the intermediate claim Lxyvw1z: SNo ((x * y + v * w1) * z).
An exact proof term for the current goal is SNo_M (x * y + v * w1) z Lxyvw1 Hz.
We prove the intermediate claim Lvyxw1z: SNo ((v * y + x * w1) * z).
An exact proof term for the current goal is SNo_M (v * y + x * w1) z Lvyxw1 Hz.
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11).
We prove the intermediate claim Lxyvw1w2: SNo ((x * y + v * w1) * w2).
An exact proof term for the current goal is SNo_M (x * y + v * w1) w2 Lxyvw1 Hw21.
We prove the intermediate claim Lvyxw1w2: SNo ((v * y + x * w1) * w2).
An exact proof term for the current goal is SNo_M (v * y + x * w1) w2 Lvyxw1 Hw21.
Apply L2 v (SNoL_SNoS x Hx v Hv) w Lw w1 (SNoL_SNoS y Hy w1 Hw1) w2 (SNoR_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove x * (w1 * z + y * w2) + v * (w + w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2).
rewrite the current goal using add_SNo_com (v * (w1 * z + y * w2)) (x * (w + w1 * w2)) Lvw1zyw2 Lxww1w2 (from left to right).
rewrite the current goal using add_SNo_com (x * (w1 * z + y * w2)) (v * (w + w1 * w2)) Lxw1zyw2 Lvww1w2 (from left to right).
We will prove v * (w + w1 * w2) + x * (w1 * z + y * w2) x * (w + w1 * w2) + v * (w1 * z + y * w2).
Apply M_Le x (w + w1 * w2) v (w1 * z + y * w2) Hx Lww1w2 Hv1 Lw1zyw2 to the current goal.
We will prove v x.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
An exact proof term for the current goal is Hwl.
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < (v * y + x * w1) * z + (x * y + v * w1) * w2.
rewrite the current goal using add_SNo_com ((v * y + x * w1) * z) ((x * y + v * w1) * w2) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using add_SNo_com ((x * y + v * w1) * z) ((v * y + x * w1) * w2) Lxyvw1z Lvyxw1w2 (from left to right).
We will prove (v * y + x * w1) * w2 + (x * y + v * w1) * z < (x * y + v * w1) * w2 + (v * y + x * w1) * z.
Apply M_Lt (x * y + v * w1) w2 (v * y + x * w1) z (SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11)) Hw21 (SNo_add_SNo (v * y) (x * w1) (SNo_M v y Hv1 Hy) (SNo_M x w1 Hx Hw11)) Hz to the current goal.
We will prove v * y + x * w1 < x * y + v * w1.
Apply M_Lt x y v w1 Hx Hy Hv1 Hw11 to the current goal.
We will prove v < x.
An exact proof term for the current goal is Hv3.
We will prove w1 < y.
An exact proof term for the current goal is Hw13.
We will prove z < w2.
An exact proof term for the current goal is Hw23.
Let w1 be given.
Assume Hw1: w1 SNoR y.
Let w2 be given.
Assume Hw2: w2 SNoL z.
Assume Hwl: w1 * z + y * w2 w + w1 * w2.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv1 Hy.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw11.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv1 Lww1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv1 Lw1zyw2.
We prove the intermediate claim Lxww1w2: SNo (x * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M x (w + w1 * w2) Hx Lww1w2.
We prove the intermediate claim Lxw1zyw2: SNo (x * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M x (w1 * z + y * w2) Hx Lw1zyw2.
Apply L2 v (SNoL_SNoS x Hx v Hv) w Lw w1 (SNoR_SNoS y Hy w1 Hw1) w2 (SNoL_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove x * (w1 * z + y * w2) + v * (w + w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2).
rewrite the current goal using add_SNo_com (v * (w1 * z + y * w2)) (x * (w + w1 * w2)) Lvw1zyw2 Lxww1w2 (from left to right).
rewrite the current goal using add_SNo_com (x * (w1 * z + y * w2)) (v * (w + w1 * w2)) Lxw1zyw2 Lvww1w2 (from left to right).
Apply M_Le x (w + w1 * w2) v (w1 * z + y * w2) Hx Lww1w2 Hv1 Lw1zyw2 to the current goal.
We will prove v x.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
An exact proof term for the current goal is Hwl.
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < (v * y + x * w1) * z + (x * y + v * w1) * w2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) (SNo_M v y Hv1 Hy) (SNo_M x w1 Hx Hw11).
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11).
Apply M_Lt (v * y + x * w1) z (x * y + v * w1) w2 Lvyxw1 Hz Lxyvw1 Hw21 to the current goal.
We will prove x * y + v * w1 < v * y + x * w1.
rewrite the current goal using add_SNo_com (x * y) (v * w1) Lxy Lvw1 (from left to right).
rewrite the current goal using add_SNo_com (v * y) (x * w1) Lvy Lxw1 (from left to right).
Apply M_Lt x w1 v y Hx Hw11 Hv1 Hy to the current goal.
We will prove v < x.
An exact proof term for the current goal is Hv3.
We will prove y < w1.
An exact proof term for the current goal is Hw13.
We will prove w2 < z.
An exact proof term for the current goal is Hw23.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoL (y * z).
Assume Hue: u = v * (y * z) + x * w + - v * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3.
We prove the intermediate claim Lw: SNo w.
Apply SNoL_E (y * z) Lyz w Hw to the current goal.
Assume H _ _.
An exact proof term for the current goal is H.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_M v w Hv1 Lw.
Apply IL y z Hy Hz w Hw to the current goal.
Let w1 be given.
Assume Hw1: w1 SNoL y.
Let w2 be given.
Assume Hw2: w2 SNoL z.
Assume Hwl: w + w1 * w2 w1 * z + y * w2.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv1 Hy.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw11.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
We prove the intermediate claim Lxww1w2: SNo (x * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M x (w + w1 * w2) Hx Lww1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv1 Lww1w2.
We prove the intermediate claim Lvw1w2: SNo (v * (w1 * w2)).
An exact proof term for the current goal is SNo_M v (w1 * w2) Hv1 Lw1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv1 Lw1zyw2.
We prove the intermediate claim Lvwvw1w2: SNo (v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo (v * w) (v * (w1 * w2)) Lvw Lvw1w2.
We prove the intermediate claim Lxw1zyw2: SNo (x * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M x (w1 * z + y * w2) Hx Lw1zyw2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) Lvy Lxw1.
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy Lvw1.
We prove the intermediate claim Lvyxw1z: SNo ((v * y + x * w1) * z).
An exact proof term for the current goal is SNo_M (v * y + x * w1) z Lvyxw1 Hz.
We prove the intermediate claim Lxyvw1z: SNo ((x * y + v * w1) * z).
An exact proof term for the current goal is SNo_M (x * y + v * w1) z Lxyvw1 Hz.
We prove the intermediate claim Lvyxw1w2: SNo ((v * y + x * w1) * w2).
An exact proof term for the current goal is SNo_M (v * y + x * w1) w2 Lvyxw1 Hw21.
We prove the intermediate claim Lxyvw1w2: SNo ((x * y + v * w1) * w2).
An exact proof term for the current goal is SNo_M (x * y + v * w1) w2 Lxyvw1 Hw21.
Apply L2 v (SNoR_SNoS x Hx v Hv) w Lw w1 (SNoL_SNoS y Hy w1 Hw1) w2 (SNoL_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove x * (w1 * z + y * w2) + v * (w + w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2).
Apply M_Le v (w1 * z + y * w2) x (w + w1 * w2) Hv1 Lw1zyw2 Hx Lww1w2 to the current goal.
We will prove x v.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
An exact proof term for the current goal is Hwl.
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < (v * y + x * w1) * z + (x * y + v * w1) * w2.
Apply M_Lt (v * y + x * w1) z (x * y + v * w1) w2 Lvyxw1 Hz Lxyvw1 Hw21 to the current goal.
We will prove x * y + v * w1 < v * y + x * w1.
Apply M_Lt v y x w1 Hv1 Hy Hx Hw11 to the current goal.
We will prove x < v.
An exact proof term for the current goal is Hv3.
We will prove w1 < y.
An exact proof term for the current goal is Hw13.
An exact proof term for the current goal is Hw23.
Let w1 be given.
Assume Hw1: w1 SNoR y.
Let w2 be given.
Assume Hw2: w2 SNoR z.
Assume Hwl: w + w1 * w2 w1 * z + y * w2.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13.
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23.
We prove the intermediate claim Lvy: SNo (v * y).
An exact proof term for the current goal is SNo_M v y Hv1 Hy.
We prove the intermediate claim Lvw1: SNo (v * w1).
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11.
We prove the intermediate claim Lxw1: SNo (x * w1).
An exact proof term for the current goal is SNo_M x w1 Hx Hw11.
We prove the intermediate claim Lw1z: SNo (w1 * z).
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz.
We prove the intermediate claim Lyw2: SNo (y * w2).
An exact proof term for the current goal is SNo_M y w2 Hy Hw21.
We prove the intermediate claim Lw1zyw2: SNo (w1 * z + y * w2).
An exact proof term for the current goal is SNo_add_SNo (w1 * z) (y * w2) Lw1z Lyw2.
We prove the intermediate claim Lw1w2: SNo (w1 * w2).
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21.
We prove the intermediate claim Lww1w2: SNo (w + w1 * w2).
An exact proof term for the current goal is SNo_add_SNo w (w1 * w2) Lw Lw1w2.
We prove the intermediate claim Lxww1w2: SNo (x * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M x (w + w1 * w2) Hx Lww1w2.
We prove the intermediate claim Lvww1w2: SNo (v * (w + w1 * w2)).
An exact proof term for the current goal is SNo_M v (w + w1 * w2) Hv1 Lww1w2.
We prove the intermediate claim Lvw1w2: SNo (v * (w1 * w2)).
An exact proof term for the current goal is SNo_M v (w1 * w2) Hv1 Lw1w2.
We prove the intermediate claim Lvw1zyw2: SNo (v * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M v (w1 * z + y * w2) Hv1 Lw1zyw2.
We prove the intermediate claim Lvwvw1w2: SNo (v * w + v * (w1 * w2)).
An exact proof term for the current goal is SNo_add_SNo (v * w) (v * (w1 * w2)) Lvw Lvw1w2.
We prove the intermediate claim Lxw1zyw2: SNo (x * (w1 * z + y * w2)).
An exact proof term for the current goal is SNo_M x (w1 * z + y * w2) Hx Lw1zyw2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) Lvy Lxw1.
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy Lvw1.
We prove the intermediate claim Lvyxw1z: SNo ((v * y + x * w1) * z).
An exact proof term for the current goal is SNo_M (v * y + x * w1) z Lvyxw1 Hz.
We prove the intermediate claim Lxyvw1z: SNo ((x * y + v * w1) * z).
An exact proof term for the current goal is SNo_M (x * y + v * w1) z Lxyvw1 Hz.
We prove the intermediate claim Lvyxw1w2: SNo ((v * y + x * w1) * w2).
An exact proof term for the current goal is SNo_M (v * y + x * w1) w2 Lvyxw1 Hw21.
We prove the intermediate claim Lxyvw1w2: SNo ((x * y + v * w1) * w2).
An exact proof term for the current goal is SNo_M (x * y + v * w1) w2 Lxyvw1 Hw21.
Apply L2 v (SNoR_SNoS x Hx v Hv) w Lw w1 (SNoR_SNoS y Hy w1 Hw1) w2 (SNoR_SNoS z Hz w2 Hw2) Hue to the current goal.
We will prove x * (w1 * z + y * w2) + v * (w + w1 * w2) v * (w1 * z + y * w2) + x * (w + w1 * w2).
Apply M_Le v (w1 * z + y * w2) x (w + w1 * w2) Hv1 Lw1zyw2 Hx Lww1w2 to the current goal.
We will prove x v.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hv3.
An exact proof term for the current goal is Hwl.
We will prove (x * y + v * w1) * z + (v * y + x * w1) * w2 < (v * y + x * w1) * z + (x * y + v * w1) * w2.
We prove the intermediate claim Lvyxw1: SNo (v * y + x * w1).
An exact proof term for the current goal is SNo_add_SNo (v * y) (x * w1) (SNo_M v y Hv1 Hy) (SNo_M x w1 Hx Hw11).
We prove the intermediate claim Lxyvw1: SNo (x * y + v * w1).
An exact proof term for the current goal is SNo_add_SNo (x * y) (v * w1) Lxy (SNo_M v w1 Hv1 Hw11).
rewrite the current goal using add_SNo_com ((v * y + x * w1) * z) ((x * y + v * w1) * w2) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using add_SNo_com ((x * y + v * w1) * z) ((v * y + x * w1) * w2) Lxyvw1z Lvyxw1w2 (from left to right).
We will prove (v * y + x * w1) * w2 + (x * y + v * w1) * z < (x * y + v * w1) * w2 + (v * y + x * w1) * z.
Apply M_Lt (x * y + v * w1) w2 (v * y + x * w1) z Lxyvw1 Hw21 Lvyxw1 Hz to the current goal.
We will prove v * y + x * w1 < x * y + v * w1.
rewrite the current goal using add_SNo_com (x * y) (v * w1) Lxy Lvw1 (from left to right).
rewrite the current goal using add_SNo_com (v * y) (x * w1) Lvy Lxw1 (from left to right).
We will prove x * w1 + v * y < v * w1 + x * y.
Apply M_Lt v w1 x y Hv1 Hw11 Hx Hy to the current goal.
We will prove x < v.
An exact proof term for the current goal is Hv3.
We will prove y < w1.
An exact proof term for the current goal is Hw13.
An exact proof term for the current goal is Hw23.
End of Section mul_SNo_assoc_lems
Theorem. (mul_SNo_assoc) The following is provable:
∀x y z, SNo xSNo ySNo zx * (y * z) = (x * y) * z
Proof:
Set P to be the term λx y z ⇒ x * (y * z) = (x * y) * z of type setsetsetprop.
We will prove ∀x y z, SNo xSNo ySNo zP x y z.
Apply SNoLev_ind3 P to the current goal.
Let x, y and z be given.
Assume Hx Hy Hz.
Assume IHx: uSNoS_ (SNoLev x), u * (y * z) = (u * y) * z.
Assume IHy: vSNoS_ (SNoLev y), x * (v * z) = (x * v) * z.
Assume IHz: wSNoS_ (SNoLev z), x * (y * w) = (x * y) * w.
Assume IHxy: uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), u * (v * z) = (u * v) * z.
Assume IHxz: uSNoS_ (SNoLev x), wSNoS_ (SNoLev z), u * (y * w) = (u * y) * w.
Assume IHyz: vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), x * (v * w) = (x * v) * w.
Assume IHxyz: uSNoS_ (SNoLev x), vSNoS_ (SNoLev y), wSNoS_ (SNoLev z), u * (v * w) = (u * v) * w.
We will prove x * (y * z) = (x * y) * z.
We prove the intermediate claim Lxy: SNo (x * y).
An exact proof term for the current goal is SNo_mul_SNo x y Hx Hy.
We prove the intermediate claim Lyz: SNo (y * z).
An exact proof term for the current goal is SNo_mul_SNo y z Hy Hz.
We prove the intermediate claim Lxyz1: SNo (x * (y * z)).
An exact proof term for the current goal is SNo_mul_SNo x (y * z) Hx Lyz.
We prove the intermediate claim Lxyz2: SNo ((x * y) * z).
An exact proof term for the current goal is SNo_mul_SNo (x * y) z Lxy Hz.
Apply mul_SNo_eq_3 x (y * z) Hx Lyz to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2.
Assume HE: x * (y * z) = SNoCut L R.
Apply mul_SNo_eq_3 (x * y) z Lxy Hz to the current goal.
Let L' and R' be given.
Assume HLR' HLE' HLI1' HLI2' HRE' HRI1' HRI2'.
Assume HE': (x * y) * z = SNoCut L' R'.
rewrite the current goal using HE (from left to right).
rewrite the current goal using HE' (from left to right).
We will prove SNoCut L R = SNoCut L' R'.
We prove the intermediate claim LIL': ∀x y, SNo xSNo yuSNoL (y * x), ∀p : prop, (vSNoL x, wSNoL y, u + w * v y * v + w * xp)(vSNoR x, wSNoR y, u + w * v y * v + w * xp)p.
Let x and y be given.
Assume Hx Hy.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp1 Hp2.
Apply mul_SNo_SNoL_interpolate_impred y x Hy Hx u Hu to the current goal.
Let w be given.
Assume Hw: w SNoL y.
Let v be given.
Assume Hv: v SNoL x.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 _ _.
We will prove u + w * v w * x + y * vp.
rewrite the current goal using add_SNo_com (w * x) (y * v) (SNo_mul_SNo w x Hw1 Hx) (SNo_mul_SNo y v Hy Hv1) (from left to right).
An exact proof term for the current goal is Hp1 v Hv w Hw.
Let w be given.
Assume Hw: w SNoR y.
Let v be given.
Assume Hv: v SNoR x.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 _ _.
We will prove u + w * v w * x + y * vp.
rewrite the current goal using add_SNo_com (w * x) (y * v) (SNo_mul_SNo w x Hw1 Hx) (SNo_mul_SNo y v Hy Hv1) (from left to right).
An exact proof term for the current goal is Hp2 v Hv w Hw.
We prove the intermediate claim LIR': ∀x y, SNo xSNo yuSNoR (y * x), ∀p : prop, (vSNoL x, wSNoR y, y * v + w * x u + w * vp)(vSNoR x, wSNoL y, y * v + w * x u + w * vp)p.
Let x and y be given.
Assume Hx Hy.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp1 Hp2.
Apply mul_SNo_SNoR_interpolate_impred y x Hy Hx u Hu to the current goal.
Let w be given.
Assume Hw: w SNoL y.
Let v be given.
Assume Hv: v SNoR x.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 _ _.
rewrite the current goal using add_SNo_com (w * x) (y * v) (SNo_mul_SNo w x Hw1 Hx) (SNo_mul_SNo y v Hy Hv1) (from left to right).
An exact proof term for the current goal is Hp2 v Hv w Hw.
Let w be given.
Assume Hw: w SNoR y.
Let v be given.
Assume Hv: v SNoL x.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 _ _.
rewrite the current goal using add_SNo_com (w * x) (y * v) (SNo_mul_SNo w x Hw1 Hx) (SNo_mul_SNo y v Hy Hv1) (from left to right).
An exact proof term for the current goal is Hp1 v Hv w Hw.
We prove the intermediate claim LMLt': ∀x y u v, SNo xSNo ySNo uSNo vu < xv < yy * u + v * x < y * x + v * u.
Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy.
rewrite the current goal using add_SNo_com (y * u) (v * x) (SNo_mul_SNo y u Hy Hu) (SNo_mul_SNo v x Hv Hx) (from left to right).
We will prove v * x + y * u < y * x + v * u.
An exact proof term for the current goal is mul_SNo_Lt y x v u Hy Hx Hv Hu Hvy Hux.
We prove the intermediate claim LMLe': ∀x y u v, SNo xSNo ySNo uSNo vu xv yy * u + v * x y * x + v * u.
Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy.
rewrite the current goal using add_SNo_com (y * u) (v * x) (SNo_mul_SNo y u Hy Hu) (SNo_mul_SNo v x Hv Hx) (from left to right).
We will prove v * x + y * u y * x + v * u.
An exact proof term for the current goal is mul_SNo_Le y x v u Hy Hx Hv Hu Hvy Hux.
We prove the intermediate claim LIHx': uSNoS_ (SNoLev x), (u * y) * z = u * (y * z).
Let u be given.
Assume Hu.
Use symmetry.
An exact proof term for the current goal is IHx u Hu.
We prove the intermediate claim LIHy': vSNoS_ (SNoLev y), (x * v) * z = x * (v * z).
Let v be given.
Assume Hv.
Use symmetry.
An exact proof term for the current goal is IHy v Hv.
We prove the intermediate claim LIHz': wSNoS_ (SNoLev z), (x * y) * w = x * (y * w).
Let w be given.
Assume Hw.
Use symmetry.
An exact proof term for the current goal is IHz w Hw.
We prove the intermediate claim LIHyx': vSNoS_ (SNoLev y), uSNoS_ (SNoLev x), (u * v) * z = u * (v * z).
Let v be given.
Assume Hv.
Let u be given.
Assume Hu.
Use symmetry.
An exact proof term for the current goal is IHxy u Hu v Hv.
We prove the intermediate claim LIHzx': wSNoS_ (SNoLev z), uSNoS_ (SNoLev x), (u * y) * w = u * (y * w).
Let w be given.
Assume Hw.
Let u be given.
Assume Hu.
Use symmetry.
An exact proof term for the current goal is IHxz u Hu w Hw.
We prove the intermediate claim LIHzy': wSNoS_ (SNoLev z), vSNoS_ (SNoLev y), (x * v) * w = x * (v * w).
Let w be given.
Assume Hw.
Let v be given.
Assume Hv.
Use symmetry.
An exact proof term for the current goal is IHyz v Hv w Hw.
We prove the intermediate claim LIHzyx': wSNoS_ (SNoLev z), vSNoS_ (SNoLev y), uSNoS_ (SNoLev x), (u * v) * w = u * (v * w).
Let w be given.
Assume Hw.
Let v be given.
Assume Hv.
Let u be given.
Assume Hu.
Use symmetry.
An exact proof term for the current goal is IHxyz u Hu v Hv w Hw.
Apply SNoCut_ext to the current goal.
An exact proof term for the current goal is HLR.
An exact proof term for the current goal is HLR'.
rewrite the current goal using HE' (from right to left).
We will prove uL, u < (x * y) * z.
An exact proof term for the current goal is mul_SNo_assoc_lem1 mul_SNo SNo_mul_SNo mul_SNo_distrL mul_SNo_distrR mul_SNo_SNoL_interpolate_impred mul_SNo_SNoR_interpolate_impred mul_SNo_Lt mul_SNo_Le x y z Hx Hy Hz IHx IHy IHz IHxy IHxz IHyz IHxyz L HLE.
rewrite the current goal using HE' (from right to left).
We will prove uR, (x * y) * z < u.
An exact proof term for the current goal is mul_SNo_assoc_lem2 mul_SNo SNo_mul_SNo mul_SNo_distrL mul_SNo_distrR mul_SNo_SNoL_interpolate_impred mul_SNo_SNoR_interpolate_impred mul_SNo_Lt mul_SNo_Le x y z Hx Hy Hz IHx IHy IHz IHxy IHxz IHyz IHxyz R HRE.
rewrite the current goal using HE (from right to left).
We will prove uL', u < x * (y * z).
Apply mul_SNo_assoc_lem1 (λx y ⇒ y * x) (λx y Hx Hy ⇒ SNo_mul_SNo y x Hy Hx) (λx y z Hx Hy Hz ⇒ mul_SNo_distrR y z x Hy Hz Hx) (λx y z Hx Hy Hz ⇒ mul_SNo_distrL z x y Hz Hx Hy) LIL' LIR' LMLt' LMLe' z y x Hz Hy Hx LIHz' LIHy' LIHx' LIHzy' LIHzx' LIHyx' LIHzyx' to the current goal.
We will prove uL', ∀q : prop, (vSNoL z, wSNoL (x * y), u = (x * y) * v + w * z + - w * vq)(vSNoR z, wSNoR (x * y), u = (x * y) * v + w * z + - w * vq)q.
Let u be given.
Assume Hu.
Let q be given.
Assume Hq1 Hq2.
Apply HLE' u Hu to the current goal.
Let w be given.
Assume Hw.
Let v be given.
Assume Hv.
Apply SNoL_E (x * y) Lxy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoL_E z Hz v Hv to the current goal.
Assume Hv1 _ _.
We will prove u = w * z + (x * y) * v + - w * vq.
rewrite the current goal using add_SNo_com_3_0_1 (w * z) ((x * y) * v) (- w * v) (SNo_mul_SNo w z Hw1 Hz) (SNo_mul_SNo (x * y) v Lxy Hv1) (SNo_minus_SNo (w * v) (SNo_mul_SNo w v Hw1 Hv1)) (from left to right).
An exact proof term for the current goal is Hq1 v Hv w Hw.
Let w be given.
Assume Hw.
Let v be given.
Assume Hv.
Apply SNoR_E (x * y) Lxy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoR_E z Hz v Hv to the current goal.
Assume Hv1 _ _.
We will prove u = w * z + (x * y) * v + - w * vq.
rewrite the current goal using add_SNo_com_3_0_1 (w * z) ((x * y) * v) (- w * v) (SNo_mul_SNo w z Hw1 Hz) (SNo_mul_SNo (x * y) v Lxy Hv1) (SNo_minus_SNo (w * v) (SNo_mul_SNo w v Hw1 Hv1)) (from left to right).
An exact proof term for the current goal is Hq2 v Hv w Hw.
rewrite the current goal using HE (from right to left).
We will prove uR', x * (y * z) < u.
Apply mul_SNo_assoc_lem2 (λx y ⇒ y * x) (λx y Hx Hy ⇒ SNo_mul_SNo y x Hy Hx) (λx y z Hx Hy Hz ⇒ mul_SNo_distrR y z x Hy Hz Hx) (λx y z Hx Hy Hz ⇒ mul_SNo_distrL z x y Hz Hx Hy) LIL' LIR' LMLt' LMLe' z y x Hz Hy Hx LIHz' LIHy' LIHx' LIHzy' LIHzx' LIHyx' LIHzyx' to the current goal.
We will prove uR', ∀q : prop, (vSNoL z, wSNoR (x * y), u = (x * y) * v + w * z + - w * vq)(vSNoR z, wSNoL (x * y), u = (x * y) * v + w * z + - w * vq)q.
Let u be given.
Assume Hu.
Let q be given.
Assume Hq1 Hq2.
Apply HRE' u Hu to the current goal.
Let w be given.
Assume Hw.
Let v be given.
Assume Hv.
Apply SNoL_E (x * y) Lxy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoR_E z Hz v Hv to the current goal.
Assume Hv1 _ _.
We will prove u = w * z + (x * y) * v + - w * vq.
rewrite the current goal using add_SNo_com_3_0_1 (w * z) ((x * y) * v) (- w * v) (SNo_mul_SNo w z Hw1 Hz) (SNo_mul_SNo (x * y) v Lxy Hv1) (SNo_minus_SNo (w * v) (SNo_mul_SNo w v Hw1 Hv1)) (from left to right).
An exact proof term for the current goal is Hq2 v Hv w Hw.
Let w be given.
Assume Hw.
Let v be given.
Assume Hv.
Apply SNoR_E (x * y) Lxy w Hw to the current goal.
Assume Hw1 _ _.
Apply SNoL_E z Hz v Hv to the current goal.
Assume Hv1 _ _.
We will prove u = w * z + (x * y) * v + - w * vq.
rewrite the current goal using add_SNo_com_3_0_1 (w * z) ((x * y) * v) (- w * v) (SNo_mul_SNo w z Hw1 Hz) (SNo_mul_SNo (x * y) v Lxy Hv1) (SNo_minus_SNo (w * v) (SNo_mul_SNo w v Hw1 Hv1)) (from left to right).
An exact proof term for the current goal is Hq1 v Hv w Hw.
Theorem. (mul_nat_mul_SNo) The following is provable:
Proof:
Let n be given.
Assume Hn: n ω.
We prove the intermediate claim Ln1: nat_p n.
An exact proof term for the current goal is omega_nat_p n Hn.
We prove the intermediate claim Ln2: ordinal n.
An exact proof term for the current goal is nat_p_ordinal n Ln1.
We prove the intermediate claim Ln3: SNo n.
An exact proof term for the current goal is ordinal_SNo n Ln2.
We prove the intermediate claim L1: ∀m, nat_p mmul_nat n m = n * m.
Apply nat_ind to the current goal.
We will prove mul_nat n 0 = n * 0.
rewrite the current goal using mul_SNo_zeroR n Ln3 (from left to right).
We will prove mul_nat n 0 = 0.
An exact proof term for the current goal is mul_nat_0R n.
Let m be given.
Assume Hm: nat_p m.
Assume IH: mul_nat n m = n * m.
We will prove mul_nat n (ordsucc m) = n * (ordsucc m).
Use transitivity with add_nat n (mul_nat n m), n + (mul_nat n m), and n + (n * m).
An exact proof term for the current goal is mul_nat_SR n m Hm.
An exact proof term for the current goal is add_nat_add_SNo n Hn (mul_nat n m) (nat_p_omega (mul_nat n m) (mul_nat_p n Ln1 m Hm)).
Use f_equal.
An exact proof term for the current goal is IH.
We will prove n + n * m = n * ordsucc m.
Use symmetry.
We will prove n * ordsucc m = n + n * m.
rewrite the current goal using add_SNo_0L m (ordinal_SNo m (nat_p_ordinal m Hm)) (from right to left) at position 1.
We will prove n * ordsucc (0 + m) = n + n * m.
rewrite the current goal using add_SNo_ordinal_SL 0 ordinal_Empty m (nat_p_ordinal m Hm) (from right to left).
We will prove n * (1 + m) = n + n * m.
rewrite the current goal using mul_SNo_distrL n 1 m Ln3 SNo_1 (ordinal_SNo m (nat_p_ordinal m Hm)) (from left to right).
We will prove n * 1 + n * m = n + n * m.
Use f_equal.
We will prove n * 1 = n.
An exact proof term for the current goal is mul_SNo_oneR n Ln3.
Let m be given.
Assume Hm: m ω.
We will prove mul_nat n m = n * m.
An exact proof term for the current goal is L1 m (omega_nat_p m Hm).
Theorem. (mul_SNo_In_omega) The following is provable:
Proof:
Let n be given.
Assume Hn.
Let m be given.
Assume Hm.
rewrite the current goal using mul_nat_mul_SNo n Hn m Hm (from right to left).
Apply nat_p_omega to the current goal.
An exact proof term for the current goal is mul_nat_p n (omega_nat_p n Hn) m (omega_nat_p m Hm).
Theorem. (mul_SNo_zeroL) The following is provable:
∀x, SNo x0 * x = 0
Proof:
Let x be given.
Assume Hx.
rewrite the current goal using mul_SNo_com 0 x SNo_0 Hx (from left to right).
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
Theorem. (mul_SNo_oneL) The following is provable:
∀x, SNo x1 * x = x
Proof:
Let x be given.
Assume Hx.
rewrite the current goal using mul_SNo_com 1 x SNo_1 Hx (from left to right).
An exact proof term for the current goal is mul_SNo_oneR x Hx.
Theorem. (SNo_gt2_double_ltS) The following is provable:
∀x, SNo x1 < xx + 1 < 2 * x
Proof:
Let x be given.
Assume Hx Hx1.
rewrite the current goal using add_SNo_1_1_2 (from right to left).
We will prove x + 1 < (1 + 1) * x.
rewrite the current goal using mul_SNo_distrR 1 1 x SNo_1 SNo_1 Hx (from left to right).
rewrite the current goal using mul_SNo_oneL x Hx (from left to right).
We will prove x + 1 < x + x.
Apply add_SNo_Lt2 x 1 x Hx SNo_1 Hx to the current goal.
We will prove 1 < x.
An exact proof term for the current goal is Hx1.
Theorem. (pos_mul_SNo_Lt) The following is provable:
∀x y z, SNo x0 < xSNo ySNo zy < zx * y < x * z
Proof:
Let x, y and z be given.
Assume Hx1 Hx2 Hy Hz Hyz.
We will prove x * y < x * z.
We prove the intermediate claim L1: 0 * z + x * y = x * y.
Use transitivity with and 0 + x * y.
Use f_equal.
We will prove 0 * z = 0.
An exact proof term for the current goal is mul_SNo_zeroL z Hz.
An exact proof term for the current goal is add_SNo_0L (x * y) (SNo_mul_SNo x y Hx1 Hy).
We prove the intermediate claim L2: x * z + 0 * y = x * z.
Use transitivity with and x * z + 0.
Use f_equal.
An exact proof term for the current goal is mul_SNo_zeroL y Hy.
An exact proof term for the current goal is add_SNo_0R (x * z) (SNo_mul_SNo x z Hx1 Hz).
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
An exact proof term for the current goal is mul_SNo_Lt x z 0 y Hx1 Hz SNo_0 Hy Hx2 Hyz.
Theorem. (nonneg_mul_SNo_Le) The following is provable:
∀x y z, SNo x0 xSNo ySNo zy zx * y x * z
Proof:
Let x, y and z be given.
Assume Hx1 Hx2 Hy Hz Hyz.
We will prove x * y x * z.
We prove the intermediate claim L1: 0 * z + x * y = x * y.
Use transitivity with and 0 + x * y.
Use f_equal.
We will prove 0 * z = 0.
An exact proof term for the current goal is mul_SNo_zeroL z Hz.
An exact proof term for the current goal is add_SNo_0L (x * y) (SNo_mul_SNo x y Hx1 Hy).
We prove the intermediate claim L2: x * z + 0 * y = x * z.
Use transitivity with and x * z + 0.
Use f_equal.
An exact proof term for the current goal is mul_SNo_zeroL y Hy.
An exact proof term for the current goal is add_SNo_0R (x * z) (SNo_mul_SNo x z Hx1 Hz).
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
An exact proof term for the current goal is mul_SNo_Le x z 0 y Hx1 Hz SNo_0 Hy Hx2 Hyz.
Theorem. (neg_mul_SNo_Lt) The following is provable:
∀x y z, SNo xx < 0SNo ySNo zz < yx * y < x * z
Proof:
Let x, y and z be given.
Assume Hx1 Hx2 Hy Hz Hzy.
We will prove x * y < x * z.
We prove the intermediate claim L1: x * y + 0 * z = x * y.
Use transitivity with and x * y + 0.
Use f_equal.
We will prove 0 * z = 0.
An exact proof term for the current goal is mul_SNo_zeroL z Hz.
An exact proof term for the current goal is add_SNo_0R (x * y) (SNo_mul_SNo x y Hx1 Hy).
We prove the intermediate claim L2: 0 * y + x * z = x * z.
Use transitivity with and 0 + x * z.
Use f_equal.
An exact proof term for the current goal is mul_SNo_zeroL y Hy.
An exact proof term for the current goal is add_SNo_0L (x * z) (SNo_mul_SNo x z Hx1 Hz).
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
An exact proof term for the current goal is mul_SNo_Lt 0 y x z SNo_0 Hy Hx1 Hz Hx2 Hzy.
Theorem. (pos_mul_SNo_Lt') The following is provable:
∀x y z, SNo xSNo ySNo z0 < zx < yx * z < y * z
Proof:
Let x, y and z be given.
Assume Hx Hy Hz Hzpos Hxy.
rewrite the current goal using mul_SNo_com x z Hx Hz (from left to right).
rewrite the current goal using mul_SNo_com y z Hy Hz (from left to right).
An exact proof term for the current goal is pos_mul_SNo_Lt z x y Hz Hzpos Hx Hy Hxy.
Theorem. (mul_SNo_Lt1_pos_Lt) The following is provable:
∀x y, SNo xSNo yx < 10 < yx * y < y
Proof:
Let x and y be given.
Assume Hx Hy Hx1 Hy0.
rewrite the current goal using mul_SNo_oneL y Hy (from right to left) at position 2.
We will prove x * y < 1 * y.
An exact proof term for the current goal is pos_mul_SNo_Lt' x 1 y Hx SNo_1 Hy Hy0 Hx1.
Theorem. (nonneg_mul_SNo_Le') The following is provable:
∀x y z, SNo xSNo ySNo z0 zx yx * z y * z
Proof:
Let x, y and z be given.
Assume Hx Hy Hz Hznn Hxy.
rewrite the current goal using mul_SNo_com x z Hx Hz (from left to right).
rewrite the current goal using mul_SNo_com y z Hy Hz (from left to right).
An exact proof term for the current goal is nonneg_mul_SNo_Le z x y Hz Hznn Hx Hy Hxy.
Theorem. (mul_SNo_Le1_nonneg_Le) The following is provable:
∀x y, SNo xSNo yx 10 yx * y y
Proof:
Let x and y be given.
Assume Hx Hy Hx1 Hy0.
rewrite the current goal using mul_SNo_oneL y Hy (from right to left) at position 2.
We will prove x * y 1 * y.
An exact proof term for the current goal is nonneg_mul_SNo_Le' x 1 y Hx SNo_1 Hy Hy0 Hx1.
Theorem. (pos_mul_SNo_Lt2) The following is provable:
∀x y z w, SNo xSNo ySNo zSNo w0 < x0 < yx < zy < wx * y < z * w
Proof:
Let x, y, z and w be given.
Assume Hx Hy Hz Hw Hxpos Hypos Hxz Hyw.
Apply SNoLt_tra (x * y) (x * w) (z * w) (SNo_mul_SNo x y Hx Hy) (SNo_mul_SNo x w Hx Hw) (SNo_mul_SNo z w Hz Hw) to the current goal.
We will prove x * y < x * w.
An exact proof term for the current goal is pos_mul_SNo_Lt x y w Hx Hxpos Hy Hw Hyw.
We will prove x * w < z * w.
Apply pos_mul_SNo_Lt' x z w Hx Hz Hw to the current goal.
We will prove 0 < w.
An exact proof term for the current goal is SNoLt_tra 0 y w SNo_0 Hy Hw Hypos Hyw.
An exact proof term for the current goal is Hxz.
Theorem. (nonneg_mul_SNo_Le2) The following is provable:
∀x y z w, SNo xSNo ySNo zSNo w0 x0 yx zy wx * y z * w
Proof:
Let x, y, z and w be given.
Assume Hx Hy Hz Hw Hxnn Hynn Hxz Hyw.
Apply SNoLe_tra (x * y) (x * w) (z * w) (SNo_mul_SNo x y Hx Hy) (SNo_mul_SNo x w Hx Hw) (SNo_mul_SNo z w Hz Hw) to the current goal.
We will prove x * y x * w.
An exact proof term for the current goal is nonneg_mul_SNo_Le x y w Hx Hxnn Hy Hw Hyw.
We will prove x * w z * w.
Apply nonneg_mul_SNo_Le' x z w Hx Hz Hw to the current goal.
We will prove 0 w.
An exact proof term for the current goal is SNoLe_tra 0 y w SNo_0 Hy Hw Hynn Hyw.
An exact proof term for the current goal is Hxz.
Theorem. (mul_SNo_pos_pos) The following is provable:
∀x y, SNo xSNo y0 < x0 < y0 < x * y
Proof:
Let x and y be given.
Assume Hx Hy Hx0 Hy0.
We will prove 0 < x * y.
rewrite the current goal using add_SNo_0R 0 SNo_0 (from right to left).
We will prove 0 + 0 < x * y.
rewrite the current goal using add_SNo_0R (x * y) (SNo_mul_SNo x y Hx Hy) (from right to left).
We will prove 0 + 0 < x * y + 0.
rewrite the current goal using mul_SNo_zeroR 0 SNo_0 (from right to left) at position 3.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left) at position 2.
rewrite the current goal using mul_SNo_zeroR y Hy (from right to left) at position 1.
We will prove y * 0 + x * 0 < x * y + 0 * 0.
rewrite the current goal using mul_SNo_com y 0 Hy SNo_0 (from left to right).
We will prove 0 * y + x * 0 < x * y + 0 * 0.
An exact proof term for the current goal is mul_SNo_Lt x y 0 0 Hx Hy SNo_0 SNo_0 Hx0 Hy0.
Theorem. (mul_SNo_pos_neg) The following is provable:
∀x y, SNo xSNo y0 < xy < 0x * y < 0
Proof:
Let x and y be given.
Assume Hx Hy Hx0 Hy0.
We will prove x * y < 0.
rewrite the current goal using add_SNo_0R 0 SNo_0 (from right to left).
rewrite the current goal using add_SNo_0L (x * y) (SNo_mul_SNo x y Hx Hy) (from right to left).
We will prove 0 + x * y < 0 + 0.
rewrite the current goal using mul_SNo_zeroR y Hy (from right to left) at position 3.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left) at position 2.
rewrite the current goal using mul_SNo_zeroR 0 SNo_0 (from right to left) at position 1.
rewrite the current goal using mul_SNo_com y 0 Hy SNo_0 (from left to right).
We will prove 0 * 0 + x * y < x * 0 + 0 * y.
An exact proof term for the current goal is mul_SNo_Lt x 0 0 y Hx SNo_0 SNo_0 Hy Hx0 Hy0.
Theorem. (mul_SNo_neg_pos) The following is provable:
∀x y, SNo xSNo yx < 00 < yx * y < 0
Proof:
Let x and y be given.
Assume Hx Hy Hx0 Hy0.
We will prove x * y < 0.
rewrite the current goal using add_SNo_0R 0 SNo_0 (from right to left).
rewrite the current goal using add_SNo_0R (x * y) (SNo_mul_SNo x y Hx Hy) (from right to left).
We will prove x * y + 0 < 0 + 0.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left) at position 3.
rewrite the current goal using mul_SNo_zeroR y Hy (from right to left) at position 2.
rewrite the current goal using mul_SNo_zeroR 0 SNo_0 (from right to left) at position 1.
rewrite the current goal using mul_SNo_com y 0 Hy SNo_0 (from left to right).
We will prove x * y + 0 * 0 < 0 * y + x * 0.
An exact proof term for the current goal is mul_SNo_Lt 0 y x 0 SNo_0 Hy Hx SNo_0 Hx0 Hy0.
Theorem. (mul_SNo_neg_neg) The following is provable:
∀x y, SNo xSNo yx < 0y < 00 < x * y
Proof:
Let x and y be given.
Assume Hx Hy Hx0 Hy0.
We will prove 0 < x * y.
rewrite the current goal using add_SNo_0R 0 SNo_0 (from right to left).
We will prove 0 + 0 < x * y.
rewrite the current goal using add_SNo_0L (x * y) (SNo_mul_SNo x y Hx Hy) (from right to left).
We will prove 0 + 0 < 0 + x * y.
rewrite the current goal using mul_SNo_zeroR 0 SNo_0 (from right to left) at position 3.
rewrite the current goal using mul_SNo_zeroR y Hy (from right to left) at position 2.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left) at position 1.
We will prove x * 0 + y * 0 < 0 * 0 + x * y.
rewrite the current goal using mul_SNo_com y 0 Hy SNo_0 (from left to right).
We will prove x * 0 + 0 * y < 0 * 0 + x * y.
An exact proof term for the current goal is mul_SNo_Lt 0 0 x y SNo_0 SNo_0 Hx Hy Hx0 Hy0.
Theorem. (mul_SNo_nonneg_nonneg) The following is provable:
∀x y, SNo xSNo y0 x0 y0 x * y
Proof:
Let x and y be given.
Assume Hx Hy Hxnn Hynn.
Apply SNoLeE 0 x SNo_0 Hx Hxnn to the current goal.
Assume H1: 0 < x.
Apply SNoLeE 0 y SNo_0 Hy Hynn to the current goal.
Assume H2: 0 < y.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is mul_SNo_pos_pos x y Hx Hy H1 H2.
Assume H2: 0 = y.
rewrite the current goal using H2 (from right to left).
rewrite the current goal using mul_SNo_zeroR x Hx (from left to right).
Apply SNoLe_ref to the current goal.
Assume H1: 0 = x.
rewrite the current goal using H1 (from right to left).
rewrite the current goal using mul_SNo_zeroL y Hy (from left to right).
Apply SNoLe_ref to the current goal.
Theorem. (mul_SNo_nonpos_pos) The following is provable:
∀x y, SNo xSNo yx 00 < yx * y 0
Proof:
Let x and y be given.
Assume Hx Hy Hxnp Hypos.
Apply SNoLtLe_or x 0 Hx SNo_0 to the current goal.
Assume H1: x < 0.
Apply SNoLtLe to the current goal.
We will prove x * y < 0.
An exact proof term for the current goal is mul_SNo_neg_pos x y Hx Hy H1 Hypos.
Assume H1: 0 x.
We prove the intermediate claim L1: x = 0.
Apply SNoLe_antisym x 0 Hx SNo_0 to the current goal.
An exact proof term for the current goal is Hxnp.
An exact proof term for the current goal is H1.
rewrite the current goal using L1 (from left to right).
We will prove 0 * y 0.
rewrite the current goal using mul_SNo_zeroL y Hy (from left to right).
Apply SNoLe_ref to the current goal.
Theorem. (mul_SNo_nonpos_neg) The following is provable:
∀x y, SNo xSNo yx 0y < 00 x * y
Proof:
Let x and y be given.
Assume Hx Hy Hxnp Hyneg.
Apply SNoLtLe_or x 0 Hx SNo_0 to the current goal.
Assume H1: x < 0.
Apply SNoLtLe to the current goal.
We will prove 0 < x * y.
An exact proof term for the current goal is mul_SNo_neg_neg x y Hx Hy H1 Hyneg.
Assume H1: 0 x.
We prove the intermediate claim L1: x = 0.
Apply SNoLe_antisym x 0 Hx SNo_0 to the current goal.
An exact proof term for the current goal is Hxnp.
An exact proof term for the current goal is H1.
rewrite the current goal using L1 (from left to right).
We will prove 0 0 * y.
rewrite the current goal using mul_SNo_zeroL y Hy (from left to right).
Apply SNoLe_ref to the current goal.
Theorem. (nonpos_mul_SNo_Le) The following is provable:
∀x y z, SNo xx 0SNo ySNo zz yx * y x * z
Proof:
Let x, y and z be given.
Assume Hx Hxnp Hy Hz Hzy.
Apply SNoLtLe_or x 0 Hx SNo_0 to the current goal.
Assume H1: x < 0.
Apply SNoLtLe_or z y Hz Hy to the current goal.
Assume H2: z < y.
Apply SNoLtLe to the current goal.
We will prove x * y < x * z.
An exact proof term for the current goal is neg_mul_SNo_Lt x y z Hx H1 Hy Hz H2.
Assume H2: y z.
We prove the intermediate claim L1: y = z.
Apply SNoLe_antisym y z Hy Hz to the current goal.
An exact proof term for the current goal is H2.
An exact proof term for the current goal is Hzy.
rewrite the current goal using L1 (from left to right).
Apply SNoLe_ref to the current goal.
Assume H1: 0 x.
We prove the intermediate claim L1: x = 0.
Apply SNoLe_antisym x 0 Hx SNo_0 to the current goal.
An exact proof term for the current goal is Hxnp.
An exact proof term for the current goal is H1.
rewrite the current goal using L1 (from left to right).
We will prove 0 * y 0 * z.
rewrite the current goal using mul_SNo_zeroL y Hy (from left to right).
rewrite the current goal using mul_SNo_zeroL z Hz (from left to right).
Apply SNoLe_ref to the current goal.
Theorem. (SNo_sqr_nonneg) The following is provable:
∀x, SNo x0 x * x
Proof:
Let x be given.
Assume Hx.
Apply SNoLtLe_or x 0 Hx SNo_0 to the current goal.
Assume H1: x < 0.
Apply SNoLtLe to the current goal.
We will prove 0 < x * x.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left).
We will prove x * 0 < x * x.
Apply neg_mul_SNo_Lt x 0 x Hx H1 SNo_0 Hx H1 to the current goal.
Assume H1: 0 x.
We will prove 0 x * x.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left).
We will prove x * 0 x * x.
An exact proof term for the current goal is nonneg_mul_SNo_Le x 0 x Hx H1 SNo_0 Hx H1.
Theorem. (SNo_zero_or_sqr_pos) The following is provable:
∀x, SNo xx = 0 0 < x * x
Proof:
Let x be given.
Assume Hx.
Apply SNoLt_trichotomy_or_impred x 0 Hx SNo_0 to the current goal.
Assume H1: x < 0.
Apply orIR to the current goal.
Apply mul_SNo_neg_neg to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is H1.
An exact proof term for the current goal is H1.
Assume H1: x = 0.
Apply orIL to the current goal.
An exact proof term for the current goal is H1.
Assume H1: 0 < x.
Apply orIR to the current goal.
Apply mul_SNo_pos_pos to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is H1.
An exact proof term for the current goal is H1.
Theorem. (SNo_pos_sqr_uniq) The following is provable:
∀x y, SNo xSNo y0 < x0 < yx * x = y * yx = y
Proof:
Let x and y be given.
Assume Hx Hy Hxpos Hypos.
Assume H1: x * x = y * y.
Apply SNoLt_trichotomy_or_impred x y Hx Hy to the current goal.
Assume H2: x < y.
We will prove False.
Apply SNoLt_irref (x * x) to the current goal.
We will prove x * x < x * x.
rewrite the current goal using H1 (from left to right) at position 2.
We will prove x * x < y * y.
An exact proof term for the current goal is pos_mul_SNo_Lt2 x x y y Hx Hx Hy Hy Hxpos Hxpos H2 H2.
Assume H2: x = y.
An exact proof term for the current goal is H2.
Assume H2: y < x.
We will prove False.
Apply SNoLt_irref (x * x) to the current goal.
We will prove x * x < x * x.
rewrite the current goal using H1 (from left to right) at position 1.
We will prove y * y < x * x.
An exact proof term for the current goal is pos_mul_SNo_Lt2 y y x x Hy Hy Hx Hx Hypos Hypos H2 H2.
Theorem. (SNo_nonneg_sqr_uniq) The following is provable:
∀x y, SNo xSNo y0 x0 yx * x = y * yx = y
Proof:
Let x and y be given.
Assume Hx Hy Hxnn Hynn.
Apply SNoLeE 0 x SNo_0 Hx Hxnn to the current goal.
Assume H1: 0 < x.
Apply SNoLeE 0 y SNo_0 Hy Hynn to the current goal.
Assume H2: 0 < y.
An exact proof term for the current goal is SNo_pos_sqr_uniq x y Hx Hy H1 H2.
Assume H2: 0 = y.
rewrite the current goal using H2 (from right to left).
rewrite the current goal using mul_SNo_zeroR 0 SNo_0 (from left to right).
Assume H3: x * x = 0.
We will prove x = 0.
Apply SNo_zero_or_sqr_pos x Hx to the current goal.
Assume H4: x = 0.
An exact proof term for the current goal is H4.
Assume H4: 0 < x * x.
We will prove False.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
rewrite the current goal using H3 (from right to left) at position 2.
An exact proof term for the current goal is H4.
Assume H1: 0 = x.
rewrite the current goal using H1 (from right to left).
rewrite the current goal using mul_SNo_zeroR 0 SNo_0 (from left to right).
Assume H2: 0 = y * y.
We will prove 0 = y.
Apply SNo_zero_or_sqr_pos y Hy to the current goal.
Assume H3: y = 0.
Use symmetry.
An exact proof term for the current goal is H3.
Assume H3: 0 < y * y.
We will prove False.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
rewrite the current goal using H2 (from left to right) at position 2.
An exact proof term for the current goal is H3.
Theorem. (SNo_foil) The following is provable:
∀x y z w, SNo xSNo ySNo zSNo w(x + y) * (z + w) = x * z + x * w + y * z + y * w
Proof:
Let x, y, z and w be given.
Assume Hx Hy Hz Hw.
Use transitivity with (x + y) * z + (x + y) * w, (x * z + y * z) + (x + y) * w, and (x * z + y * z) + (x * w + y * w).
An exact proof term for the current goal is mul_SNo_distrL (x + y) z w (SNo_add_SNo x y Hx Hy) Hz Hw.
Use f_equal.
An exact proof term for the current goal is mul_SNo_distrR x y z Hx Hy Hz.
Use f_equal.
An exact proof term for the current goal is mul_SNo_distrR x y w Hx Hy Hw.
We will prove (x * z + y * z) + (x * w + y * w) = x * z + (x * w + (y * z + y * w)).
rewrite the current goal using add_SNo_com_4_inner_mid (x * z) (y * z) (x * w) (y * w) (SNo_mul_SNo x z Hx Hz) (SNo_mul_SNo y z Hy Hz) (SNo_mul_SNo x w Hx Hw) (SNo_mul_SNo y w Hy Hw) (from left to right).
We will prove (x * z + x * w) + (y * z + y * w) = x * z + (x * w + (y * z + y * w)).
Use symmetry.
An exact proof term for the current goal is add_SNo_assoc (x * z) (x * w) (y * z + y * w) (SNo_mul_SNo x z Hx Hz) (SNo_mul_SNo x w Hx Hw) (SNo_add_SNo (y * z) (y * w) (SNo_mul_SNo y z Hy Hz) (SNo_mul_SNo y w Hy Hw)).
Theorem. (mul_SNo_minus_minus) The following is provable:
∀x y, SNo xSNo y(- x) * (- y) = x * y
Proof:
Let x and y be given.
Assume Hx Hy.
rewrite the current goal using mul_SNo_minus_distrL x (- y) Hx (SNo_minus_SNo y Hy) (from left to right).
We will prove - (x * (- y)) = x * y.
rewrite the current goal using mul_SNo_minus_distrR x y Hx Hy (from left to right).
We will prove - - (x * y) = x * y.
An exact proof term for the current goal is minus_SNo_invol (x * y) (SNo_mul_SNo x y Hx Hy).
Theorem. (mul_SNo_com_3_0_1) The following is provable:
∀x y z, SNo xSNo ySNo zx * y * z = y * x * z
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
rewrite the current goal using mul_SNo_assoc x y z Hx Hy Hz (from left to right).
rewrite the current goal using mul_SNo_assoc y x z Hy Hx Hz (from left to right).
Use f_equal.
An exact proof term for the current goal is mul_SNo_com x y Hx Hy.
Theorem. (mul_SNo_com_3b_1_2) The following is provable:
∀x y z, SNo xSNo ySNo z(x * y) * z = (x * z) * y
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
rewrite the current goal using mul_SNo_assoc x y z Hx Hy Hz (from right to left).
rewrite the current goal using mul_SNo_assoc x z y Hx Hz Hy (from right to left).
Use f_equal.
An exact proof term for the current goal is mul_SNo_com y z Hy Hz.
Theorem. (mul_SNo_com_4_inner_mid) The following is provable:
∀x y z w, SNo xSNo ySNo zSNo w(x * y) * (z * w) = (x * z) * (y * w)
Proof:
Let x, y, z and w be given.
Assume Hx Hy Hz Hw.
rewrite the current goal using mul_SNo_assoc (x * y) z w (SNo_mul_SNo x y Hx Hy) Hz Hw (from left to right).
We will prove ((x * y) * z) * w = (x * z) * (y * w).
rewrite the current goal using mul_SNo_com_3b_1_2 x y z Hx Hy Hz (from left to right).
We will prove ((x * z) * y) * w = (x * z) * (y * w).
Use symmetry.
An exact proof term for the current goal is mul_SNo_assoc (x * z) y w (SNo_mul_SNo x z Hx Hz) Hy Hw.
Theorem. (mul_SNo_rotate_3_1) The following is provable:
∀x y z, SNo xSNo ySNo zx * y * z = z * x * y
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
We will prove x * (y * z) = z * (x * y).
Use transitivity with x * (z * y), (x * z) * y, and (z * x) * y.
Use f_equal.
An exact proof term for the current goal is mul_SNo_com y z Hy Hz.
An exact proof term for the current goal is mul_SNo_assoc x z y Hx Hz Hy.
Use f_equal.
An exact proof term for the current goal is mul_SNo_com x z Hx Hz.
Use symmetry.
An exact proof term for the current goal is mul_SNo_assoc z x y Hz Hx Hy.
Theorem. (mul_SNo_rotate_4_1) The following is provable:
∀x y z w, SNo xSNo ySNo zSNo wx * y * z * w = w * x * y * z
Proof:
Let x, y, z and w be given.
Assume Hx Hy Hz Hw.
rewrite the current goal using mul_SNo_rotate_3_1 y z w Hy Hz Hw (from left to right).
We will prove x * w * y * z = w * x * y * z.
An exact proof term for the current goal is mul_SNo_com_3_0_1 x w (y * z) Hx Hw (SNo_mul_SNo y z Hy Hz).
Theorem. (SNo_foil_mm) The following is provable:
∀x y z w, SNo xSNo ySNo zSNo w(x + - y) * (z + - w) = x * z + - x * w + - y * z + y * w
Proof:
Let x, y, z and w be given.
Assume Hx Hy Hz Hw.
We prove the intermediate claim Lmy: SNo (- y).
An exact proof term for the current goal is SNo_minus_SNo y Hy.
We prove the intermediate claim Lmw: SNo (- w).
An exact proof term for the current goal is SNo_minus_SNo w Hw.
rewrite the current goal using SNo_foil x (- y) z (- w) Hx Lmy Hz Lmw (from left to right).
We will prove x * z + x * (- w) + (- y) * z + (- y) * (- w) = x * z + - x * w + - y * z + y * w.
rewrite the current goal using mul_SNo_minus_minus y w Hy Hw (from left to right).
rewrite the current goal using mul_SNo_minus_distrL y z Hy Hz (from left to right).
rewrite the current goal using mul_SNo_minus_distrR x w Hx Hw (from left to right).
Use reflexivity.
Theorem. (mul_SNo_nonzero_cancel) The following is provable:
∀x y z, SNo xx 0SNo ySNo zx * y = x * zy = z
Proof:
Let x, y and z be given.
Assume Hx Hxn0 Hy Hz Hxyz.
Apply SNoLt_trichotomy_or_impred y z Hy Hz to the current goal.
Assume H1: y < z.
We will prove False.
Apply SNoLt_trichotomy_or_impred x 0 Hx SNo_0 to the current goal.
Assume H2: x < 0.
Apply SNoLt_irref (x * y) to the current goal.
We will prove x * y < x * y.
rewrite the current goal using Hxyz (from left to right) at position 1.
We will prove x * z < x * y.
An exact proof term for the current goal is neg_mul_SNo_Lt x z y Hx H2 Hz Hy H1.
Assume H2: x = 0.
An exact proof term for the current goal is Hxn0 H2.
Assume H2: 0 < x.
Apply SNoLt_irref (x * y) to the current goal.
We will prove x * y < x * y.
rewrite the current goal using Hxyz (from left to right) at position 2.
We will prove x * y < x * z.
An exact proof term for the current goal is pos_mul_SNo_Lt x y z Hx H2 Hy Hz H1.
Assume H1.
An exact proof term for the current goal is H1.
Assume H1: z < y.
We will prove False.
Apply SNoLt_trichotomy_or_impred x 0 Hx SNo_0 to the current goal.
Assume H2: x < 0.
Apply SNoLt_irref (x * y) to the current goal.
We will prove x * y < x * y.
rewrite the current goal using Hxyz (from left to right) at position 2.
We will prove x * y < x * z.
An exact proof term for the current goal is neg_mul_SNo_Lt x y z Hx H2 Hy Hz H1.
Assume H2: x = 0.
An exact proof term for the current goal is Hxn0 H2.
Assume H2: 0 < x.
Apply SNoLt_irref (x * y) to the current goal.
We will prove x * y < x * y.
rewrite the current goal using Hxyz (from left to right) at position 1.
We will prove x * z < x * y.
An exact proof term for the current goal is pos_mul_SNo_Lt x z y Hx H2 Hz Hy H1.
Theorem. (mul_SNoCutP_lem) The following is provable:
∀Lx Rx Ly Ry x y, SNoCutP Lx RxSNoCutP Ly Ryx = SNoCut Lx Rxy = SNoCut Ly RySNoCutP ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ly} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ry}) ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ry} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ly}) x * y = SNoCut ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ly} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ry}) ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ry} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ly}) ∀q : prop, (∀LxLy' RxRy' LxRy' RxLy', (uLxLy', ∀p : prop, (wLx, w'Ly, SNo wSNo w'w < xw' < yu = w * y + x * w' + - w * w'p)p)(wLx, w'Ly, w * y + x * w' + - w * w' LxLy')(uRxRy', ∀p : prop, (zRx, z'Ry, SNo zSNo z'x < zy < z'u = z * y + x * z' + - z * z'p)p)(zRx, z'Ry, z * y + x * z' + - z * z' RxRy')(uLxRy', ∀p : prop, (wLx, zRy, SNo wSNo zw < xy < zu = w * y + x * z + - w * zp)p)(wLx, zRy, w * y + x * z + - w * z LxRy')(uRxLy', ∀p : prop, (zRx, wLy, SNo zSNo wx < zw < yu = z * y + x * w + - z * wp)p)(zRx, wLy, z * y + x * w + - z * w RxLy')SNoCutP (LxLy' RxRy') (LxRy' RxLy')x * y = SNoCut (LxLy' RxRy') (LxRy' RxLy')q)q
Proof:
Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye.
Apply HLRx to the current goal.
Assume H.
Apply H to the current goal.
Assume HLRx1: wLx, SNo w.
Assume HLRx2: zRx, SNo z.
Assume HLRx3: wLx, zRx, w < z.
Apply HLRy to the current goal.
Assume H.
Apply H to the current goal.
Assume HLRy1: wLy, SNo w.
Assume HLRy2: zRy, SNo z.
Assume HLRy3: wLy, zRy, w < z.
Apply SNoCutP_SNoCut_impred Lx Rx HLRx to the current goal.
rewrite the current goal using Hxe (from right to left).
Assume Hx1: SNo x.
Assume Hx2: SNoLev x ordsucc ((wLxordsucc (SNoLev w)) (zRxordsucc (SNoLev z))).
Assume Hx3: wLx, w < x.
Assume Hx4: zRx, x < z.
Assume Hx5: (∀u, SNo u(wLx, w < u)(zRx, u < z)SNoLev x SNoLev u SNoEq_ (SNoLev x) x u).
Apply SNoCutP_SNoCut_impred Ly Ry HLRy to the current goal.
rewrite the current goal using Hye (from right to left).
Assume Hy1: SNo y.
Assume Hy2: SNoLev y ordsucc ((wLyordsucc (SNoLev w)) (zRyordsucc (SNoLev z))).
Assume Hy3: wLy, w < y.
Assume Hy4: zRy, y < z.
Assume Hy5: (∀u, SNo u(wLy, w < u)(zRy, u < z)SNoLev y SNoLev u SNoEq_ (SNoLev y) y u).
Set LxLy' to be the term {w 0 * y + x * w 1 + - w 0 * w 1|wLx Ly}.
Set RxRy' to be the term {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ry}.
Set LxRy' to be the term {w 0 * y + x * w 1 + - w 0 * w 1|wLx Ry}.
Set RxLy' to be the term {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ly}.
We prove the intermediate claim LxLy'E: uLxLy', ∀p : prop, (wLx, w'Ly, SNo wSNo w'w < xw' < yu = w * y + x * w' + - w * w'p)p.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp.
Apply ReplE_impred (Lx Ly) (λw ⇒ (w 0) * y + x * (w 1) + - (w 0) * (w 1)) u Hu to the current goal.
Let ww' be given.
Assume Hww': ww' Lx Ly.
Assume Hue: u = (ww' 0) * y + x * (ww' 1) + - (ww' 0) * (ww' 1).
We prove the intermediate claim Lww'0: ww' 0 Lx.
An exact proof term for the current goal is ap0_Sigma Lx (λ_ ⇒ Ly) ww' Hww'.
We prove the intermediate claim Lww'1: ww' 1 Ly.
An exact proof term for the current goal is ap1_Sigma Lx (λ_ ⇒ Ly) ww' Hww'.
Apply Hp (ww' 0) Lww'0 (ww' 1) Lww'1 to the current goal.
We will prove SNo (ww' 0).
An exact proof term for the current goal is HLRx1 (ww' 0) Lww'0.
We will prove SNo (ww' 1).
An exact proof term for the current goal is HLRy1 (ww' 1) Lww'1.
We will prove ww' 0 < x.
An exact proof term for the current goal is Hx3 (ww' 0) Lww'0.
We will prove ww' 1 < y.
An exact proof term for the current goal is Hy3 (ww' 1) Lww'1.
An exact proof term for the current goal is Hue.
We prove the intermediate claim LxLy'I: wLx, w'Ly, w * y + x * w' + - w * w' LxLy'.
Let w be given.
Assume Hw.
Let w' be given.
Assume Hw'.
rewrite the current goal using tuple_2_0_eq w w' (from right to left).
rewrite the current goal using tuple_2_1_eq w w' (from right to left) at positions 2 and 4.
Apply ReplI (Lx Ly) (λw ⇒ w 0 * y + x * w 1 + - w 0 * w 1) (w,w') to the current goal.
We will prove (w,w') Lx Ly.
An exact proof term for the current goal is tuple_2_setprod Lx Ly w Hw w' Hw'.
We prove the intermediate claim RxRy'E: uRxRy', ∀p : prop, (zRx, z'Ry, SNo zSNo z'x < zy < z'u = z * y + x * z' + - z * z'p)p.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp.
Apply ReplE_impred (Rx Ry) (λz ⇒ (z 0) * y + x * (z 1) + - (z 0) * (z 1)) u Hu to the current goal.
Let zz' be given.
Assume Hzz': zz' Rx Ry.
Assume Hue: u = (zz' 0) * y + x * (zz' 1) + - (zz' 0) * (zz' 1).
We prove the intermediate claim Lzz'0: zz' 0 Rx.
An exact proof term for the current goal is ap0_Sigma Rx (λ_ ⇒ Ry) zz' Hzz'.
We prove the intermediate claim Lzz'1: zz' 1 Ry.
An exact proof term for the current goal is ap1_Sigma Rx (λ_ ⇒ Ry) zz' Hzz'.
Apply Hp (zz' 0) Lzz'0 (zz' 1) Lzz'1 to the current goal.
We will prove SNo (zz' 0).
An exact proof term for the current goal is HLRx2 (zz' 0) Lzz'0.
We will prove SNo (zz' 1).
An exact proof term for the current goal is HLRy2 (zz' 1) Lzz'1.
We will prove x < zz' 0.
An exact proof term for the current goal is Hx4 (zz' 0) Lzz'0.
We will prove y < zz' 1.
An exact proof term for the current goal is Hy4 (zz' 1) Lzz'1.
An exact proof term for the current goal is Hue.
We prove the intermediate claim RxRy'I: zRx, z'Ry, z * y + x * z' + - z * z' RxRy'.
Let z be given.
Assume Hz.
Let z' be given.
Assume Hz'.
rewrite the current goal using tuple_2_0_eq z z' (from right to left).
rewrite the current goal using tuple_2_1_eq z z' (from right to left) at positions 2 and 4.
Apply ReplI (Rx Ry) (λz ⇒ z 0 * y + x * z 1 + - z 0 * z 1) (z,z') to the current goal.
We will prove (z,z') Rx Ry.
An exact proof term for the current goal is tuple_2_setprod Rx Ry z Hz z' Hz'.
We prove the intermediate claim LxRy'E: uLxRy', ∀p : prop, (wLx, zRy, SNo wSNo zw < xy < zu = w * y + x * z + - w * zp)p.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp.
Apply ReplE_impred (Lx Ry) (λw ⇒ (w 0) * y + x * (w 1) + - (w 0) * (w 1)) u Hu to the current goal.
Let wz be given.
Assume Hwz: wz Lx Ry.
Assume Hue: u = (wz 0) * y + x * (wz 1) + - (wz 0) * (wz 1).
We prove the intermediate claim Lwz0: wz 0 Lx.
An exact proof term for the current goal is ap0_Sigma Lx (λ_ ⇒ Ry) wz Hwz.
We prove the intermediate claim Lwz1: wz 1 Ry.
An exact proof term for the current goal is ap1_Sigma Lx (λ_ ⇒ Ry) wz Hwz.
Apply Hp (wz 0) Lwz0 (wz 1) Lwz1 to the current goal.
We will prove SNo (wz 0).
An exact proof term for the current goal is HLRx1 (wz 0) Lwz0.
We will prove SNo (wz 1).
An exact proof term for the current goal is HLRy2 (wz 1) Lwz1.
We will prove wz 0 < x.
An exact proof term for the current goal is Hx3 (wz 0) Lwz0.
We will prove y < wz 1.
An exact proof term for the current goal is Hy4 (wz 1) Lwz1.
An exact proof term for the current goal is Hue.
We prove the intermediate claim LxRy'I: wLx, zRy, w * y + x * z + - w * z LxRy'.
Let w be given.
Assume Hw.
Let z be given.
Assume Hz.
rewrite the current goal using tuple_2_0_eq w z (from right to left).
rewrite the current goal using tuple_2_1_eq w z (from right to left) at positions 2 and 4.
Apply ReplI (Lx Ry) (λw ⇒ w 0 * y + x * w 1 + - w 0 * w 1) (w,z) to the current goal.
We will prove (w,z) Lx Ry.
An exact proof term for the current goal is tuple_2_setprod Lx Ry w Hw z Hz.
We prove the intermediate claim RxLy'E: uRxLy', ∀p : prop, (zRx, wLy, SNo zSNo wx < zw < yu = z * y + x * w + - z * wp)p.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp.
Apply ReplE_impred (Rx Ly) (λz ⇒ (z 0) * y + x * (z 1) + - (z 0) * (z 1)) u Hu to the current goal.
Let zw be given.
Assume Hzw: zw Rx Ly.
Assume Hue: u = (zw 0) * y + x * (zw 1) + - (zw 0) * (zw 1).
We prove the intermediate claim Lzw0: zw 0 Rx.
An exact proof term for the current goal is ap0_Sigma Rx (λ_ ⇒ Ly) zw Hzw.
We prove the intermediate claim Lzw1: zw 1 Ly.
An exact proof term for the current goal is ap1_Sigma Rx (λ_ ⇒ Ly) zw Hzw.
Apply Hp (zw 0) Lzw0 (zw 1) Lzw1 to the current goal.
We will prove SNo (zw 0).
An exact proof term for the current goal is HLRx2 (zw 0) Lzw0.
We will prove SNo (zw 1).
An exact proof term for the current goal is HLRy1 (zw 1) Lzw1.
We will prove x < zw 0.
An exact proof term for the current goal is Hx4 (zw 0) Lzw0.
We will prove zw 1 < y.
An exact proof term for the current goal is Hy3 (zw 1) Lzw1.
An exact proof term for the current goal is Hue.
We prove the intermediate claim RxLy'I: zRx, wLy, z * y + x * w + - z * w RxLy'.
Let z be given.
Assume Hz.
Let w be given.
Assume Hw.
rewrite the current goal using tuple_2_0_eq z w (from right to left).
rewrite the current goal using tuple_2_1_eq z w (from right to left) at positions 2 and 4.
Apply ReplI (Rx Ly) (λw ⇒ w 0 * y + x * w 1 + - w 0 * w 1) (z,w) to the current goal.
We will prove (z,w) Rx Ly.
An exact proof term for the current goal is tuple_2_setprod Rx Ly z Hz w Hw.
We prove the intermediate claim L1: SNoCutP (LxLy' RxRy') (LxRy' RxLy').
We will prove (wLxLy' RxRy', SNo w) (zLxRy' RxLy', SNo z) (wLxLy' RxRy', zLxRy' RxLy', w < z).
Apply and3I to the current goal.
Let w be given.
Apply binunionE' to the current goal.
Assume Hw: w LxLy'.
Apply LxLy'E w Hw to the current goal.
Let w' be given.
Assume Hw': w' Lx.
Let w'' be given.
Assume Hw'': w'' Ly.
Assume Hw'1: SNo w'.
Assume Hw''1: SNo w''.
Assume Hw'2: w' < x.
Assume Hw''2: w'' < y.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
We will prove SNo (w' * y + x * w'' + - w' * w'').
Apply SNo_add_SNo_3 to the current goal.
An exact proof term for the current goal is SNo_mul_SNo w' y Hw'1 Hy1.
An exact proof term for the current goal is SNo_mul_SNo x w'' Hx1 Hw''1.
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is SNo_mul_SNo w' w'' Hw'1 Hw''1.
Assume Hw: w RxRy'.
Apply RxRy'E w Hw to the current goal.
Let z' be given.
Assume Hz': z' Rx.
Let z'' be given.
Assume Hz'': z'' Ry.
Assume Hz'1: SNo z'.
Assume Hz''1: SNo z''.
Assume Hz'2: x < z'.
Assume Hz''2: y < z''.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
We will prove SNo (z' * y + x * z'' + - z' * z'').
Apply SNo_add_SNo_3 to the current goal.
An exact proof term for the current goal is SNo_mul_SNo z' y Hz'1 Hy1.
An exact proof term for the current goal is SNo_mul_SNo x z'' Hx1 Hz''1.
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is SNo_mul_SNo z' z'' Hz'1 Hz''1.
Let z be given.
Apply binunionE' to the current goal.
Assume Hz: z LxRy'.
Apply LxRy'E z Hz to the current goal.
Let w' be given.
Assume Hw': w' Lx.
Let z' be given.
Assume Hz': z' Ry.
Assume Hw'1: SNo w'.
Assume Hz'1: SNo z'.
Assume Hw'2: w' < x.
Assume Hz'2: y < z'.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove SNo (w' * y + x * z' + - w' * z').
Apply SNo_add_SNo_3 to the current goal.
An exact proof term for the current goal is SNo_mul_SNo w' y Hw'1 Hy1.
An exact proof term for the current goal is SNo_mul_SNo x z' Hx1 Hz'1.
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is SNo_mul_SNo w' z' Hw'1 Hz'1.
Assume Hz: z RxLy'.
Apply RxLy'E z Hz to the current goal.
Let z' be given.
Assume Hz': z' Rx.
Let w' be given.
Assume Hw': w' Ly.
Assume Hz'1: SNo z'.
Assume Hw'1: SNo w'.
Assume Hz'2: x < z'.
Assume Hw'2: w' < y.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove SNo (z' * y + x * w' + - z' * w').
Apply SNo_add_SNo_3 to the current goal.
An exact proof term for the current goal is SNo_mul_SNo z' y Hz'1 Hy1.
An exact proof term for the current goal is SNo_mul_SNo x w' Hx1 Hw'1.
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is SNo_mul_SNo z' w' Hz'1 Hw'1.
Let w be given.
Apply binunionE' to the current goal.
Assume Hw: w LxLy'.
Apply LxLy'E w Hw to the current goal.
Let w' be given.
Assume Hw': w' Lx.
Let w'' be given.
Assume Hw'': w'' Ly.
Assume Hw'1: SNo w'.
Assume Hw''1: SNo w''.
Assume Hw'2: w' < x.
Assume Hw''2: w'' < y.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
Let z be given.
Apply binunionE' to the current goal.
Assume Hz: z LxRy'.
Apply LxRy'E z Hz to the current goal.
Let w''' be given.
Assume Hw''': w''' Lx.
Let z' be given.
Assume Hz': z' Ry.
Assume Hw'''1: SNo w'''.
Assume Hz'1: SNo z'.
Assume Hw'''2: w''' < x.
Assume Hz'2: y < z'.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove w' * y + x * w'' + - w' * w'' < w''' * y + x * z' + - w''' * z'.
Apply add_SNo_minus_Lt12b3 (w' * y) (x * w'') (w' * w'') (w''' * y) (x * z') (w''' * z') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_mul_SNo w''' y Hw'''1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1) to the current goal.
We will prove w' * y + x * w'' + w''' * z' < w''' * y + x * z' + w' * w''.
Apply SNoLt_tra (w' * y + x * w'' + w''' * z') (x * y + w' * w'' + w''' * z') (w''' * y + x * z' + w' * w'') (SNo_add_SNo_3 (w' * y) (x * w'') (w''' * z') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1)) (SNo_add_SNo_3 (x * y) (w' * w'') (w''' * z') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1)) (SNo_add_SNo_3 (w''' * y) (x * z') (w' * w'') (SNo_mul_SNo w''' y Hw'''1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (SNo_mul_SNo w' w'' Hw'1 Hw''1)) to the current goal.
We will prove w' * y + x * w'' + w''' * z' < x * y + w' * w'' + w''' * z'.
rewrite the current goal using add_SNo_rotate_3_1 (w' * y) (x * w'') (w''' * z') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1) (from left to right).
rewrite the current goal using add_SNo_rotate_3_1 (x * y) (w' * w'') (w''' * z') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1) (from left to right).
We will prove w''' * z' + w' * y + x * w'' < w''' * z' + x * y + w' * w''.
Apply add_SNo_Lt2 (w''' * z') (w' * y + x * w'') (x * y + w' * w'') (SNo_mul_SNo w''' z' Hw'''1 Hz'1) (SNo_add_SNo (w' * y) (x * w'') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1)) (SNo_add_SNo (x * y) (w' * w'') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w' w'' Hw'1 Hw''1)) to the current goal.
We will prove w' * y + x * w'' < x * y + w' * w''.
An exact proof term for the current goal is mul_SNo_Lt x y w' w'' Hx1 Hy1 Hw'1 Hw''1 Hw'2 Hw''2.
We will prove x * y + w' * w'' + w''' * z' < w''' * y + x * z' + w' * w''.
rewrite the current goal using add_SNo_com_3_0_1 (x * y) (w' * w'') (w''' * z') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1) (from left to right).
rewrite the current goal using add_SNo_rotate_3_1 (w''' * y) (x * z') (w' * w'') (SNo_mul_SNo w''' y Hw'''1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (from left to right).
We will prove w' * w'' + x * y + w''' * z' < w' * w'' + w''' * y + x * z'.
Apply add_SNo_Lt2 (w' * w'') (x * y + w''' * z') (w''' * y + x * z') (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_add_SNo (x * y) (w''' * z') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1)) (SNo_add_SNo (w''' * y) (x * z') (SNo_mul_SNo w''' y Hw'''1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1)) to the current goal.
We will prove x * y + w''' * z' < w''' * y + x * z'.
rewrite the current goal using add_SNo_com (x * y) (w''' * z') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w''' z' Hw'''1 Hz'1) (from left to right).
rewrite the current goal using add_SNo_com (w''' * y) (x * z') (SNo_mul_SNo w''' y Hw'''1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (from left to right).
We will prove w''' * z' + x * y < x * z' + w''' * y.
An exact proof term for the current goal is mul_SNo_Lt x z' w''' y Hx1 Hz'1 Hw'''1 Hy1 Hw'''2 Hz'2.
Assume Hz: z RxLy'.
Apply RxLy'E z Hz to the current goal.
Let z' be given.
Assume Hz': z' Rx.
Let w''' be given.
Assume Hw''': w''' Ly.
Assume Hz'1: SNo z'.
Assume Hw'''1: SNo w'''.
Assume Hz'2: x < z'.
Assume Hw'''2: w''' < y.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove w' * y + x * w'' + - w' * w'' < z' * y + x * w''' + - z' * w'''.
Apply add_SNo_minus_Lt12b3 (w' * y) (x * w'') (w' * w'') (z' * y) (x * w''') (z' * w''') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x w''' Hx1 Hw'''1) (SNo_mul_SNo z' w''' Hz'1 Hw'''1) to the current goal.
We will prove w' * y + x * w'' + z' * w''' < z' * y + x * w''' + w' * w''.
Apply SNoLt_tra (w' * y + x * w'' + z' * w''') (z' * w''' + x * y + w' * w'') (z' * y + x * w''' + w' * w'') (SNo_add_SNo_3 (w' * y) (x * w'') (z' * w''') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1) (SNo_mul_SNo z' w''' Hz'1 Hw'''1)) (SNo_add_SNo_3 (z' * w''') (x * y) (w' * w'') (SNo_mul_SNo z' w''' Hz'1 Hw'''1) (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w' w'' Hw'1 Hw''1)) (SNo_add_SNo_3 (z' * y) (x * w''') (w' * w'') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x w''' Hx1 Hw'''1) (SNo_mul_SNo w' w'' Hw'1 Hw''1)) to the current goal.
We will prove w' * y + x * w'' + z' * w''' < z' * w''' + x * y + w' * w''.
rewrite the current goal using add_SNo_rotate_3_1 (w' * y) (x * w'') (z' * w''') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1) (SNo_mul_SNo z' w''' Hz'1 Hw'''1) (from left to right).
We will prove z' * w''' + w' * y + x * w'' < z' * w''' + x * y + w' * w''.
Apply add_SNo_Lt2 (z' * w''') (w' * y + x * w'') (x * y + w' * w'') (SNo_mul_SNo z' w''' Hz'1 Hw'''1) (SNo_add_SNo (w' * y) (x * w'') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1)) (SNo_add_SNo (x * y) (w' * w'') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w' w'' Hw'1 Hw''1)) to the current goal.
We will prove w' * y + x * w'' < x * y + w' * w''.
An exact proof term for the current goal is mul_SNo_Lt x y w' w'' Hx1 Hy1 Hw'1 Hw''1 Hw'2 Hw''2.
We will prove z' * w''' + x * y + w' * w'' < z' * y + x * w''' + w' * w''.
rewrite the current goal using add_SNo_rotate_3_1 (z' * w''') (x * y) (w' * w'') (SNo_mul_SNo z' w''' Hz'1 Hw'''1) (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (from left to right).
rewrite the current goal using add_SNo_rotate_3_1 (z' * y) (x * w''') (w' * w'') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x w''' Hx1 Hw'''1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (from left to right).
We will prove w' * w'' + z' * w''' + x * y < w' * w'' + z' * y + x * w'''.
Apply add_SNo_Lt2 (w' * w'') (z' * w''' + x * y) (z' * y + x * w''') (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_add_SNo (z' * w''') (x * y) (SNo_mul_SNo z' w''' Hz'1 Hw'''1) (SNo_mul_SNo x y Hx1 Hy1)) (SNo_add_SNo (z' * y) (x * w''') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x w''' Hx1 Hw'''1)) to the current goal.
We will prove z' * w''' + x * y < z' * y + x * w'''.
rewrite the current goal using add_SNo_com (z' * w''') (x * y) (SNo_mul_SNo z' w''' Hz'1 Hw'''1) (SNo_mul_SNo x y Hx1 Hy1) (from left to right).
We will prove x * y + z' * w''' < z' * y + x * w'''.
An exact proof term for the current goal is mul_SNo_Lt z' y x w''' Hz'1 Hy1 Hx1 Hw'''1 Hz'2 Hw'''2.
Assume Hw: w RxRy'.
Apply RxRy'E w Hw to the current goal.
Let z' be given.
Assume Hz': z' Rx.
Let z'' be given.
Assume Hz'': z'' Ry.
Assume Hz'1: SNo z'.
Assume Hz''1: SNo z''.
Assume Hz'2: x < z'.
Assume Hz''2: y < z''.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
Let z be given.
Apply binunionE' to the current goal.
Assume Hz: z LxRy'.
Apply LxRy'E z Hz to the current goal.
Let w' be given.
Assume Hw': w' Lx.
Let z''' be given.
Assume Hz''': z''' Ry.
Assume Hw'1: SNo w'.
Assume Hz'''1: SNo z'''.
Assume Hw'2: w' < x.
Assume Hz'''2: y < z'''.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove z' * y + x * z'' + - z' * z'' < w' * y + x * z''' + - w' * z'''.
Apply add_SNo_minus_Lt12b3 (z' * y) (x * z'') (z' * z'') (w' * y) (x * z''') (w' * z''') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x z''' Hx1 Hz'''1) (SNo_mul_SNo w' z''' Hw'1 Hz'''1) to the current goal.
We will prove z' * y + x * z'' + w' * z''' < w' * y + x * z''' + z' * z''.
Apply SNoLt_tra (z' * y + x * z'' + w' * z''') (w' * z''' + z' * z'' + x * y) (w' * y + x * z''' + z' * z'') (SNo_add_SNo_3 (z' * y) (x * z'') (w' * z''') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (SNo_mul_SNo w' z''' Hw'1 Hz'''1)) (SNo_add_SNo_3 (w' * z''') (z' * z'') (x * y) (SNo_mul_SNo w' z''' Hw'1 Hz'''1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo x y Hx1 Hy1)) (SNo_add_SNo_3 (w' * y) (x * z''') (z' * z'') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x z''' Hx1 Hz'''1) (SNo_mul_SNo z' z'' Hz'1 Hz''1)) to the current goal.
We will prove z' * y + x * z'' + w' * z''' < w' * z''' + z' * z'' + x * y.
rewrite the current goal using add_SNo_rotate_3_1 (z' * y) (x * z'') (w' * z''') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (SNo_mul_SNo w' z''' Hw'1 Hz'''1) (from left to right).
We will prove w' * z''' + z' * y + x * z'' < w' * z''' + z' * z'' + x * y.
Apply add_SNo_Lt2 (w' * z''') (z' * y + x * z'') (z' * z'' + x * y) (SNo_mul_SNo w' z''' Hw'1 Hz'''1) (SNo_add_SNo (z' * y) (x * z'') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1)) (SNo_add_SNo (z' * z'') (x * y) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo x y Hx1 Hy1)) to the current goal.
We will prove z' * y + x * z'' < z' * z'' + x * y.
rewrite the current goal using add_SNo_com (z' * y) (x * z'') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (from left to right).
An exact proof term for the current goal is mul_SNo_Lt z' z'' x y Hz'1 Hz''1 Hx1 Hy1 Hz'2 Hz''2.
We will prove w' * z''' + z' * z'' + x * y < w' * y + x * z''' + z' * z''.
rewrite the current goal using add_SNo_com_3_0_1 (w' * z''') (z' * z'') (x * y) (SNo_mul_SNo w' z''' Hw'1 Hz'''1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo x y Hx1 Hy1) (from left to right).
rewrite the current goal using add_SNo_rotate_3_1 (w' * y) (x * z''') (z' * z'') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x z''' Hx1 Hz'''1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (from left to right).
We will prove z' * z'' + w' * z''' + x * y < z' * z'' + w' * y + x * z'''.
Apply add_SNo_Lt2 (z' * z'') (w' * z''' + x * y) (w' * y + x * z''') (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_add_SNo (w' * z''') (x * y) (SNo_mul_SNo w' z''' Hw'1 Hz'''1) (SNo_mul_SNo x y Hx1 Hy1)) (SNo_add_SNo (w' * y) (x * z''') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x z''' Hx1 Hz'''1)) to the current goal.
We will prove w' * z''' + x * y < w' * y + x * z'''.
rewrite the current goal using add_SNo_com (w' * y) (x * z''') (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x z''' Hx1 Hz'''1) (from left to right).
An exact proof term for the current goal is mul_SNo_Lt x z''' w' y Hx1 Hz'''1 Hw'1 Hy1 Hw'2 Hz'''2.
Assume Hz: z RxLy'.
Apply RxLy'E z Hz to the current goal.
Let z''' be given.
Assume Hz''': z''' Rx.
Let w' be given.
Assume Hw': w' Ly.
Assume Hz'''1: SNo z'''.
Assume Hw'1: SNo w'.
Assume Hz'''2: x < z'''.
Assume Hw'2: w' < y.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove z' * y + x * z'' + - z' * z'' < z''' * y + x * w' + - z''' * w'.
Apply add_SNo_minus_Lt12b3 (z' * y) (x * z'') (z' * z'') (z''' * y) (x * w') (z''' * w') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo z''' y Hz'''1 Hy1) (SNo_mul_SNo x w' Hx1 Hw'1) (SNo_mul_SNo z''' w' Hz'''1 Hw'1) to the current goal.
We will prove z' * y + x * z'' + z''' * w' < z''' * y + x * w' + z' * z''.
Apply SNoLt_tra (z' * y + x * z'' + z''' * w') (z''' * w' + z' * z'' + x * y) (z''' * y + x * w' + z' * z'') (SNo_add_SNo_3 (z' * y) (x * z'') (z''' * w') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (SNo_mul_SNo z''' w' Hz'''1 Hw'1)) (SNo_add_SNo_3 (z''' * w') (z' * z'') (x * y) (SNo_mul_SNo z''' w' Hz'''1 Hw'1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo x y Hx1 Hy1)) (SNo_add_SNo_3 (z''' * y) (x * w') (z' * z'') (SNo_mul_SNo z''' y Hz'''1 Hy1) (SNo_mul_SNo x w' Hx1 Hw'1) (SNo_mul_SNo z' z'' Hz'1 Hz''1)) to the current goal.
We will prove z' * y + x * z'' + z''' * w' < z''' * w' + z' * z'' + x * y.
rewrite the current goal using add_SNo_rotate_3_1 (z' * y) (x * z'') (z''' * w') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (SNo_mul_SNo z''' w' Hz'''1 Hw'1) (from left to right).
rewrite the current goal using add_SNo_com (z' * y) (x * z'') (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x z'' Hx1 Hz''1) (from left to right).
We will prove z''' * w' + x * z'' + z' * y < z''' * w' + z' * z'' + x * y.
Apply add_SNo_Lt2 (z''' * w') (x * z'' + z' * y) (z' * z'' + x * y) (SNo_mul_SNo z''' w' Hz'''1 Hw'1) (SNo_add_SNo (x * z'') (z' * y) (SNo_mul_SNo x z'' Hx1 Hz''1) (SNo_mul_SNo z' y Hz'1 Hy1)) (SNo_add_SNo (z' * z'') (x * y) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo x y Hx1 Hy1)) to the current goal.
We will prove x * z'' + z' * y < z' * z'' + x * y.
An exact proof term for the current goal is mul_SNo_Lt z' z'' x y Hz'1 Hz''1 Hx1 Hy1 Hz'2 Hz''2.
We will prove z''' * w' + z' * z'' + x * y < z''' * y + x * w' + z' * z''.
rewrite the current goal using add_SNo_com_3_0_1 (z''' * w') (z' * z'') (x * y) (SNo_mul_SNo z''' w' Hz'''1 Hw'1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_mul_SNo x y Hx1 Hy1) (from left to right).
rewrite the current goal using add_SNo_rotate_3_1 (z''' * y) (x * w') (z' * z'') (SNo_mul_SNo z''' y Hz'''1 Hy1) (SNo_mul_SNo x w' Hx1 Hw'1) (SNo_mul_SNo z' z'' Hz'1 Hz''1) (from left to right).
We will prove z' * z'' + z''' * w' + x * y < z' * z'' + z''' * y + x * w'.
Apply add_SNo_Lt2 (z' * z'') (z''' * w' + x * y) (z''' * y + x * w') (SNo_mul_SNo z' z'' Hz'1 Hz''1) (SNo_add_SNo (z''' * w') (x * y) (SNo_mul_SNo z''' w' Hz'''1 Hw'1) (SNo_mul_SNo x y Hx1 Hy1)) (SNo_add_SNo (z''' * y) (x * w') (SNo_mul_SNo z''' y Hz'''1 Hy1) (SNo_mul_SNo x w' Hx1 Hw'1)) to the current goal.
We will prove z''' * w' + x * y < z''' * y + x * w'.
rewrite the current goal using add_SNo_com (z''' * w') (x * y) (SNo_mul_SNo z''' w' Hz'''1 Hw'1) (SNo_mul_SNo x y Hx1 Hy1) (from left to right).
An exact proof term for the current goal is mul_SNo_Lt z''' y x w' Hz'''1 Hy1 Hx1 Hw'1 Hz'''2 Hw'2.
We prove the intermediate claim Lxyeq: x * y = SNoCut (LxLy' RxRy') (LxRy' RxLy').
Set v to be the term SNoCut (LxLy' RxRy') (LxRy' RxLy').
Apply SNoCutP_SNoCut_impred (LxLy' RxRy') (LxRy' RxLy') L1 to the current goal.
Assume Hv1: SNo v.
Assume Hv2: SNoLev v ordsucc ((wLxLy' RxRy'ordsucc (SNoLev w)) (zLxRy' RxLy'ordsucc (SNoLev z))).
Assume Hv3: wLxLy' RxRy', w < v.
Assume Hv4: zLxRy' RxLy', v < z.
Assume Hv5: ∀u, SNo u(wLxLy' RxRy', w < u)(zLxRy' RxLy', u < z)SNoLev v SNoLev u SNoEq_ (SNoLev v) v u.
Apply mul_SNo_eq_3 x y Hx1 Hy1 to the current goal.
Let L' and R' be given.
Assume HLR': SNoCutP L' R'.
Assume HL'E: ∀u, u L'∀q : prop, (w0SNoL x, w1SNoL y, u = w0 * y + x * w1 + - w0 * w1q)(z0SNoR x, z1SNoR y, u = z0 * y + x * z1 + - z0 * z1q)q.
Assume HL'I1: w0SNoL x, w1SNoL y, w0 * y + x * w1 + - w0 * w1 L'.
Assume HL'I2: z0SNoR x, z1SNoR y, z0 * y + x * z1 + - z0 * z1 L'.
Assume HR'E: ∀u, u R'∀q : prop, (w0SNoL x, z1SNoR y, u = w0 * y + x * z1 + - w0 * z1q)(z0SNoR x, w1SNoL y, u = z0 * y + x * w1 + - z0 * w1q)q.
Assume HR'I1: w0SNoL x, z1SNoR y, w0 * y + x * z1 + - w0 * z1 R'.
Assume HR'I2: z0SNoR x, w1SNoL y, z0 * y + x * w1 + - z0 * w1 R'.
Assume HLR'eq: x * y = SNoCut L' R'.
rewrite the current goal using HLR'eq (from left to right).
Apply SNoCut_ext L' R' (LxLy' RxRy') (LxRy' RxLy') to the current goal.
An exact proof term for the current goal is HLR'.
An exact proof term for the current goal is L1.
We will prove wL', w < v.
Let w be given.
Assume Hw.
Apply HL'E w Hw to the current goal.
Let w0 be given.
Assume Hw0: w0 SNoL x.
Let w1 be given.
Assume Hw1: w1 SNoL y.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
We will prove w0 * y + x * w1 + - w0 * w1 < v.
Apply SNoL_E x Hx1 w0 Hw0 to the current goal.
Assume Hw0a Hw0b Hw0c.
Apply SNoL_E y Hy1 w1 Hw1 to the current goal.
Assume Hw1a Hw1b Hw1c.
We prove the intermediate claim L2: w0'Lx, w0 w0'.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L2a: SNoLev x SNoLev w0 SNoEq_ (SNoLev x) x w0.
Apply Hx5 w0 Hw0a to the current goal.
We will prove w'Lx, w' < w0.
Let w' be given.
Assume Hw'.
Apply SNoLtLe_or w' w0 (HLRx1 w' Hw') Hw0a to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: w0 w'.
We will prove False.
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw'.
We will prove w0 w'.
An exact proof term for the current goal is H.
We will prove z'Rx, w0 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLt_tra w0 x z' Hw0a Hx1 (HLRx2 z' Hz') to the current goal.
We will prove w0 < x.
An exact proof term for the current goal is Hw0c.
We will prove x < z'.
An exact proof term for the current goal is Hx4 z' Hz'.
Apply L2a to the current goal.
Assume Hxw0: SNoLev x SNoLev w0.
Assume _.
Apply In_irref (SNoLev w0) to the current goal.
We will prove SNoLev w0 SNoLev w0.
Apply Hxw0 to the current goal.
An exact proof term for the current goal is Hw0b.
We prove the intermediate claim L3: w1'Ly, w1 w1'.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L3a: SNoLev y SNoLev w1 SNoEq_ (SNoLev y) y w1.
Apply Hy5 w1 Hw1a to the current goal.
We will prove w'Ly, w' < w1.
Let w' be given.
Assume Hw'.
Apply SNoLtLe_or w' w1 (HLRy1 w' Hw') Hw1a to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: w1 w'.
We will prove False.
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw'.
We will prove w1 w'.
An exact proof term for the current goal is H.
We will prove z'Ry, w1 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLt_tra w1 y z' Hw1a Hy1 (HLRy2 z' Hz') to the current goal.
We will prove w1 < y.
An exact proof term for the current goal is Hw1c.
We will prove y < z'.
An exact proof term for the current goal is Hy4 z' Hz'.
Apply L3a to the current goal.
Assume Hyw1: SNoLev y SNoLev w1.
Assume _.
Apply In_irref (SNoLev w1) to the current goal.
We will prove SNoLev w1 SNoLev w1.
Apply Hyw1 to the current goal.
An exact proof term for the current goal is Hw1b.
Apply L2 to the current goal.
Let w0' be given.
Assume H.
Apply H to the current goal.
Assume Hw0'1: w0' Lx.
Assume Hw0'2: w0 w0'.
Apply L3 to the current goal.
Let w1' be given.
Assume H.
Apply H to the current goal.
Assume Hw1'1: w1' Ly.
Assume Hw1'2: w1 w1'.
We will prove w0 * y + x * w1 + - w0 * w1 < v.
Apply SNoLeLt_tra (w0 * y + x * w1 + - w0 * w1) (w0' * y + x * w1' + - w0' * w1') v (SNo_add_SNo_3 (w0 * y) (x * w1) (- w0 * w1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_minus_SNo (w0 * w1) (SNo_mul_SNo w0 w1 Hw0a Hw1a))) (SNo_add_SNo_3 (w0' * y) (x * w1') (- w0' * w1') (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_minus_SNo (w0' * w1') (SNo_mul_SNo w0' w1' (HLRx1 w0' Hw0'1) (HLRy1 w1' Hw1'1)))) Hv1 to the current goal.
We will prove w0 * y + x * w1 + - w0 * w1 w0' * y + x * w1' + - w0' * w1'.
Apply add_SNo_minus_Le12b3 (w0 * y) (x * w1) (w0 * w1) (w0' * y) (x * w1') (w0' * w1') (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo w0 w1 Hw0a Hw1a) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo w0' w1' (HLRx1 w0' Hw0'1) (HLRy1 w1' Hw1'1)) to the current goal.
We will prove w0 * y + x * w1 + w0' * w1' w0' * y + x * w1' + w0 * w1.
Apply SNoLe_tra (w0 * y + x * w1 + w0' * w1') (w0 * y + x * w1' + w0' * w1) (w0' * y + x * w1' + w0 * w1) (SNo_add_SNo_3 (w0 * y) (x * w1) (w0' * w1') (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo w0' w1' (HLRx1 w0' Hw0'1) (HLRy1 w1' Hw1'1))) (SNo_add_SNo_3 (w0 * y) (x * w1') (w0' * w1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo w0' w1 (HLRx1 w0' Hw0'1) Hw1a)) (SNo_add_SNo_3 (w0' * y) (x * w1') (w0 * w1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo w0 w1 Hw0a Hw1a)) to the current goal.
We will prove w0 * y + x * w1 + w0' * w1' w0 * y + x * w1' + w0' * w1.
Apply add_SNo_Le2 (w0 * y) (x * w1 + w0' * w1') (x * w1' + w0' * w1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_add_SNo (x * w1) (w0' * w1') (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo w0' w1' (HLRx1 w0' Hw0'1) (HLRy1 w1' Hw1'1))) (SNo_add_SNo (x * w1') (w0' * w1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo w0' w1 (HLRx1 w0' Hw0'1) Hw1a)) to the current goal.
We will prove x * w1 + w0' * w1' x * w1' + w0' * w1.
rewrite the current goal using add_SNo_com (x * w1) (w0' * w1') (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo w0' w1' (HLRx1 w0' Hw0'1) (HLRy1 w1' Hw1'1)) (from left to right).
We will prove w0' * w1' + x * w1 x * w1' + w0' * w1.
Apply mul_SNo_Le x w1' w0' w1 Hx1 (HLRy1 w1' Hw1'1) (HLRx1 w0' Hw0'1) Hw1a to the current goal.
We will prove w0' x.
Apply SNoLtLe to the current goal.
We will prove w0' < x.
An exact proof term for the current goal is Hx3 w0' Hw0'1.
We will prove w1 w1'.
An exact proof term for the current goal is Hw1'2.
We will prove w0 * y + x * w1' + w0' * w1 w0' * y + x * w1' + w0 * w1.
rewrite the current goal using add_SNo_com_3_0_1 (w0 * y) (x * w1') (w0' * w1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo w0' w1 (HLRx1 w0' Hw0'1) Hw1a) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (w0' * y) (x * w1') (w0 * w1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo w0 w1 Hw0a Hw1a) (from left to right).
Apply add_SNo_Le2 (x * w1') (w0 * y + w0' * w1) (w0' * y + w0 * w1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_add_SNo (w0 * y) (w0' * w1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo w0' w1 (HLRx1 w0' Hw0'1) Hw1a)) (SNo_add_SNo (w0' * y) (w0 * w1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo w0 w1 Hw0a Hw1a)) to the current goal.
We will prove w0 * y + w0' * w1 w0' * y + w0 * w1.
Apply mul_SNo_Le w0' y w0 w1 (HLRx1 w0' Hw0'1) Hy1 Hw0a Hw1a to the current goal.
We will prove w0 w0'.
An exact proof term for the current goal is Hw0'2.
We will prove w1 y.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hw1c.
We will prove w0' * y + x * w1' + - w0' * w1' < v.
Apply Hv3 to the current goal.
We will prove w0' * y + x * w1' + - w0' * w1' LxLy' RxRy'.
Apply binunionI1 to the current goal.
Apply LxLy'I to the current goal.
An exact proof term for the current goal is Hw0'1.
An exact proof term for the current goal is Hw1'1.
Let z0 be given.
Assume Hz0: z0 SNoR x.
Let z1 be given.
Assume Hz1: z1 SNoR y.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
We will prove z0 * y + x * z1 + - z0 * z1 < v.
Apply SNoR_E x Hx1 z0 Hz0 to the current goal.
Assume Hz0a Hz0b Hz0c.
Apply SNoR_E y Hy1 z1 Hz1 to the current goal.
Assume Hz1a Hz1b Hz1c.
We prove the intermediate claim L4: z0'Rx, z0' z0.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L4a: SNoLev x SNoLev z0 SNoEq_ (SNoLev x) x z0.
Apply Hx5 z0 Hz0a to the current goal.
We will prove w'Lx, w' < z0.
Let w' be given.
Assume Hw'.
Apply SNoLt_tra w' x z0 (HLRx1 w' Hw') Hx1 Hz0a to the current goal.
We will prove w' < x.
An exact proof term for the current goal is Hx3 w' Hw'.
We will prove x < z0.
An exact proof term for the current goal is Hz0c.
We will prove z'Rx, z0 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLtLe_or z0 z' Hz0a (HLRx2 z' Hz') to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: z' z0.
We will prove False.
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz'.
We will prove z' z0.
An exact proof term for the current goal is H.
Apply L4a to the current goal.
Assume Hxz0: SNoLev x SNoLev z0.
Assume _.
Apply In_irref (SNoLev z0) to the current goal.
We will prove SNoLev z0 SNoLev z0.
Apply Hxz0 to the current goal.
An exact proof term for the current goal is Hz0b.
We prove the intermediate claim L5: z1'Ry, z1' z1.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L5a: SNoLev y SNoLev z1 SNoEq_ (SNoLev y) y z1.
Apply Hy5 z1 Hz1a to the current goal.
We will prove w'Ly, w' < z1.
Let w' be given.
Assume Hw'.
Apply SNoLt_tra w' y z1 (HLRy1 w' Hw') Hy1 Hz1a to the current goal.
We will prove w' < y.
An exact proof term for the current goal is Hy3 w' Hw'.
We will prove y < z1.
An exact proof term for the current goal is Hz1c.
We will prove z'Ry, z1 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLtLe_or z1 z' Hz1a (HLRy2 z' Hz') to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: z' z1.
We will prove False.
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz'.
We will prove z' z1.
An exact proof term for the current goal is H.
Apply L5a to the current goal.
Assume Hyz1: SNoLev y SNoLev z1.
Assume _.
Apply In_irref (SNoLev z1) to the current goal.
We will prove SNoLev z1 SNoLev z1.
Apply Hyz1 to the current goal.
An exact proof term for the current goal is Hz1b.
Apply L4 to the current goal.
Let z0' be given.
Assume H.
Apply H to the current goal.
Assume Hz0'1: z0' Rx.
Assume Hz0'2: z0' z0.
Apply L5 to the current goal.
Let z1' be given.
Assume H.
Apply H to the current goal.
Assume Hz1'1: z1' Ry.
Assume Hz1'2: z1' z1.
We will prove z0 * y + x * z1 + - z0 * z1 < v.
Apply SNoLeLt_tra (z0 * y + x * z1 + - z0 * z1) (z0' * y + x * z1' + - z0' * z1') v (SNo_add_SNo_3 (z0 * y) (x * z1) (- z0 * z1) (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_minus_SNo (z0 * z1) (SNo_mul_SNo z0 z1 Hz0a Hz1a))) (SNo_add_SNo_3 (z0' * y) (x * z1') (- z0' * z1') (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_minus_SNo (z0' * z1') (SNo_mul_SNo z0' z1' (HLRx2 z0' Hz0'1) (HLRy2 z1' Hz1'1)))) Hv1 to the current goal.
We will prove z0 * y + x * z1 + - z0 * z1 z0' * y + x * z1' + - z0' * z1'.
Apply add_SNo_minus_Le12b3 (z0 * y) (x * z1) (z0 * z1) (z0' * y) (x * z1') (z0' * z1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo z0 z1 Hz0a Hz1a) (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo z0' z1' (HLRx2 z0' Hz0'1) (HLRy2 z1' Hz1'1)) to the current goal.
We will prove z0 * y + x * z1 + z0' * z1' z0' * y + x * z1' + z0 * z1.
Apply SNoLe_tra (z0 * y + x * z1 + z0' * z1') (z0 * y + z0' * z1 + x * z1') (z0' * y + x * z1' + z0 * z1) (SNo_add_SNo_3 (z0 * y) (x * z1) (z0' * z1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo z0' z1' (HLRx2 z0' Hz0'1) (HLRy2 z1' Hz1'1))) (SNo_add_SNo_3 (z0 * y) (z0' * z1) (x * z1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo z0' z1 (HLRx2 z0' Hz0'1) Hz1a) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1))) (SNo_add_SNo_3 (z0' * y) (x * z1') (z0 * z1) (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo z0 z1 Hz0a Hz1a)) to the current goal.
We will prove z0 * y + x * z1 + z0' * z1' z0 * y + z0' * z1 + x * z1'.
Apply add_SNo_Le2 (z0 * y) (x * z1 + z0' * z1') (z0' * z1 + x * z1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_add_SNo (x * z1) (z0' * z1') (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo z0' z1' (HLRx2 z0' Hz0'1) (HLRy2 z1' Hz1'1))) (SNo_add_SNo (z0' * z1) (x * z1') (SNo_mul_SNo z0' z1 (HLRx2 z0' Hz0'1) Hz1a) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1))) to the current goal.
We will prove x * z1 + z0' * z1' z0' * z1 + x * z1'.
Apply mul_SNo_Le z0' z1 x z1' (HLRx2 z0' Hz0'1) Hz1a Hx1 (HLRy2 z1' Hz1'1) to the current goal.
We will prove x z0'.
Apply SNoLtLe to the current goal.
We will prove x < z0'.
An exact proof term for the current goal is Hx4 z0' Hz0'1.
We will prove z1' z1.
An exact proof term for the current goal is Hz1'2.
We will prove z0 * y + z0' * z1 + x * z1' z0' * y + x * z1' + z0 * z1.
rewrite the current goal using add_SNo_rotate_3_1 (z0 * y) (z0' * z1) (x * z1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo z0' z1 (HLRx2 z0' Hz0'1) Hz1a) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (z0' * y) (x * z1') (z0 * z1) (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo z0 z1 Hz0a Hz1a) (from left to right).
We will prove x * z1' + z0 * y + z0' * z1 x * z1' + z0' * y + z0 * z1.
Apply add_SNo_Le2 (x * z1') (z0 * y + z0' * z1) (z0' * y + z0 * z1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_add_SNo (z0 * y) (z0' * z1) (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo z0' z1 (HLRx2 z0' Hz0'1) Hz1a)) (SNo_add_SNo (z0' * y) (z0 * z1) (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo z0 z1 Hz0a Hz1a)) to the current goal.
We will prove z0 * y + z0' * z1 z0' * y + z0 * z1.
rewrite the current goal using add_SNo_com (z0 * y) (z0' * z1) (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo z0' z1 (HLRx2 z0' Hz0'1) Hz1a) (from left to right).
rewrite the current goal using add_SNo_com (z0' * y) (z0 * z1) (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo z0 z1 Hz0a Hz1a) (from left to right).
Apply mul_SNo_Le z0 z1 z0' y Hz0a Hz1a (HLRx2 z0' Hz0'1) Hy1 to the current goal.
We will prove z0' z0.
An exact proof term for the current goal is Hz0'2.
We will prove y z1.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hz1c.
We will prove z0' * y + x * z1' + - z0' * z1' < v.
Apply Hv3 to the current goal.
We will prove z0' * y + x * z1' + - z0' * z1' LxLy' RxRy'.
Apply binunionI2 to the current goal.
Apply RxRy'I to the current goal.
An exact proof term for the current goal is Hz0'1.
An exact proof term for the current goal is Hz1'1.
We will prove zR', v < z.
Let z be given.
Assume Hz.
Apply HR'E z Hz to the current goal.
Let w0 be given.
Assume Hw0: w0 SNoL x.
Let z1 be given.
Assume Hz1: z1 SNoR y.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove v < w0 * y + x * z1 + - w0 * z1.
Apply SNoL_E x Hx1 w0 Hw0 to the current goal.
Assume Hw0a Hw0b Hw0c.
Apply SNoR_E y Hy1 z1 Hz1 to the current goal.
Assume Hz1a Hz1b Hz1c.
We prove the intermediate claim L6: w0'Lx, w0 w0'.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L6a: SNoLev x SNoLev w0 SNoEq_ (SNoLev x) x w0.
Apply Hx5 w0 Hw0a to the current goal.
We will prove w'Lx, w' < w0.
Let w' be given.
Assume Hw'.
Apply SNoLtLe_or w' w0 (HLRx1 w' Hw') Hw0a to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: w0 w'.
We will prove False.
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw'.
We will prove w0 w'.
An exact proof term for the current goal is H.
We will prove z'Rx, w0 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLt_tra w0 x z' Hw0a Hx1 (HLRx2 z' Hz') to the current goal.
We will prove w0 < x.
An exact proof term for the current goal is Hw0c.
We will prove x < z'.
An exact proof term for the current goal is Hx4 z' Hz'.
Apply L6a to the current goal.
Assume Hxw0: SNoLev x SNoLev w0.
Assume _.
Apply In_irref (SNoLev w0) to the current goal.
We will prove SNoLev w0 SNoLev w0.
Apply Hxw0 to the current goal.
An exact proof term for the current goal is Hw0b.
We prove the intermediate claim L7: z1'Ry, z1' z1.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L7a: SNoLev y SNoLev z1 SNoEq_ (SNoLev y) y z1.
Apply Hy5 z1 Hz1a to the current goal.
We will prove w'Ly, w' < z1.
Let w' be given.
Assume Hw'.
Apply SNoLt_tra w' y z1 (HLRy1 w' Hw') Hy1 Hz1a to the current goal.
We will prove w' < y.
An exact proof term for the current goal is Hy3 w' Hw'.
We will prove y < z1.
An exact proof term for the current goal is Hz1c.
We will prove z'Ry, z1 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLtLe_or z1 z' Hz1a (HLRy2 z' Hz') to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: z' z1.
We will prove False.
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz'.
We will prove z' z1.
An exact proof term for the current goal is H.
Apply L7a to the current goal.
Assume Hyz1: SNoLev y SNoLev z1.
Assume _.
Apply In_irref (SNoLev z1) to the current goal.
We will prove SNoLev z1 SNoLev z1.
Apply Hyz1 to the current goal.
An exact proof term for the current goal is Hz1b.
Apply L6 to the current goal.
Let w0' be given.
Assume H.
Apply H to the current goal.
Assume Hw0'1: w0' Lx.
Assume Hw0'2: w0 w0'.
Apply L7 to the current goal.
Let z1' be given.
Assume H.
Apply H to the current goal.
Assume Hz1'1: z1' Ry.
Assume Hz1'2: z1' z1.
We will prove v < w0 * y + x * z1 + - w0 * z1.
Apply SNoLtLe_tra v (w0' * y + x * z1' + - w0' * z1') (w0 * y + x * z1 + - w0 * z1) Hv1 (SNo_add_SNo_3 (w0' * y) (x * z1') (- w0' * z1') (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_minus_SNo (w0' * z1') (SNo_mul_SNo w0' z1' (HLRx1 w0' Hw0'1) (HLRy2 z1' Hz1'1)))) (SNo_add_SNo_3 (w0 * y) (x * z1) (- w0 * z1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_minus_SNo (w0 * z1) (SNo_mul_SNo w0 z1 Hw0a Hz1a))) to the current goal.
We will prove v < w0' * y + x * z1' + - w0' * z1'.
Apply Hv4 to the current goal.
We will prove w0' * y + x * z1' + - w0' * z1' LxRy' RxLy'.
Apply binunionI1 to the current goal.
Apply LxRy'I to the current goal.
An exact proof term for the current goal is Hw0'1.
An exact proof term for the current goal is Hz1'1.
We will prove w0' * y + x * z1' + - w0' * z1' w0 * y + x * z1 + - w0 * z1.
Apply add_SNo_minus_Le12b3 (w0' * y) (x * z1') (w0' * z1') (w0 * y) (x * z1) (w0 * z1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0' z1' (HLRx1 w0' Hw0'1) (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo w0 z1 Hw0a Hz1a) to the current goal.
We will prove w0' * y + x * z1' + w0 * z1 w0 * y + x * z1 + w0' * z1'.
Apply SNoLe_tra (w0' * y + x * z1' + w0 * z1) (x * z1' + w0' * z1 + w0 * y) (w0 * y + x * z1 + w0' * z1') (SNo_add_SNo_3 (w0' * y) (x * z1') (w0 * z1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0 z1 Hw0a Hz1a)) (SNo_add_SNo_3 (x * z1') (w0' * z1) (w0 * y) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0' z1 (HLRx1 w0' Hw0'1) Hz1a) (SNo_mul_SNo w0 y Hw0a Hy1)) (SNo_add_SNo_3 (w0 * y) (x * z1) (w0' * z1') (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo w0' z1' (HLRx1 w0' Hw0'1) (HLRy2 z1' Hz1'1))) to the current goal.
We will prove w0' * y + x * z1' + w0 * z1 x * z1' + w0' * z1 + w0 * y.
rewrite the current goal using add_SNo_com_3_0_1 (w0' * y) (x * z1') (w0 * z1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0 z1 Hw0a Hz1a) (from left to right).
Apply add_SNo_Le2 (x * z1') (w0' * y + w0 * z1) (w0' * z1 + w0 * y) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_add_SNo (w0' * y) (w0 * z1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo w0 z1 Hw0a Hz1a)) (SNo_add_SNo (w0' * z1) (w0 * y) (SNo_mul_SNo w0' z1 (HLRx1 w0' Hw0'1) Hz1a) (SNo_mul_SNo w0 y Hw0a Hy1)) to the current goal.
We will prove w0' * y + w0 * z1 w0' * z1 + w0 * y.
rewrite the current goal using add_SNo_com (w0' * y) (w0 * z1) (SNo_mul_SNo w0' y (HLRx1 w0' Hw0'1) Hy1) (SNo_mul_SNo w0 z1 Hw0a Hz1a) (from left to right).
Apply mul_SNo_Le w0' z1 w0 y (HLRx1 w0' Hw0'1) Hz1a Hw0a Hy1 to the current goal.
We will prove w0 w0'.
An exact proof term for the current goal is Hw0'2.
We will prove y z1.
Apply SNoLtLe to the current goal.
We will prove y < z1.
An exact proof term for the current goal is Hz1c.
We will prove x * z1' + w0' * z1 + w0 * y w0 * y + x * z1 + w0' * z1'.
rewrite the current goal using add_SNo_rotate_3_1 (x * z1') (w0' * z1) (w0 * y) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0' z1 (HLRx1 w0' Hw0'1) Hz1a) (SNo_mul_SNo w0 y Hw0a Hy1) (from left to right).
We will prove w0 * y + x * z1' + w0' * z1 w0 * y + x * z1 + w0' * z1'.
Apply add_SNo_Le2 (w0 * y) (x * z1' + w0' * z1) (x * z1 + w0' * z1') (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_add_SNo (x * z1') (w0' * z1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0' z1 (HLRx1 w0' Hw0'1) Hz1a)) (SNo_add_SNo (x * z1) (w0' * z1') (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo w0' z1' (HLRx1 w0' Hw0'1) (HLRy2 z1' Hz1'1))) to the current goal.
We will prove x * z1' + w0' * z1 x * z1 + w0' * z1'.
rewrite the current goal using add_SNo_com (x * z1') (w0' * z1) (SNo_mul_SNo x z1' Hx1 (HLRy2 z1' Hz1'1)) (SNo_mul_SNo w0' z1 (HLRx1 w0' Hw0'1) Hz1a) (from left to right).
We will prove w0' * z1 + x * z1' x * z1 + w0' * z1'.
Apply mul_SNo_Le x z1 w0' z1' Hx1 Hz1a (HLRx1 w0' Hw0'1) (HLRy2 z1' Hz1'1) to the current goal.
We will prove w0' x.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hx3 w0' Hw0'1.
We will prove z1' z1.
An exact proof term for the current goal is Hz1'2.
Let z0 be given.
Assume Hz0: z0 SNoR x.
Let w1 be given.
Assume Hw1: w1 SNoL y.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove v < z0 * y + x * w1 + - z0 * w1.
Apply SNoR_E x Hx1 z0 Hz0 to the current goal.
Assume Hz0a Hz0b Hz0c.
Apply SNoL_E y Hy1 w1 Hw1 to the current goal.
Assume Hw1a Hw1b Hw1c.
We prove the intermediate claim L8: z0'Rx, z0' z0.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L8a: SNoLev x SNoLev z0 SNoEq_ (SNoLev x) x z0.
Apply Hx5 z0 Hz0a to the current goal.
We will prove w'Lx, w' < z0.
Let w' be given.
Assume Hw'.
Apply SNoLt_tra w' x z0 (HLRx1 w' Hw') Hx1 Hz0a to the current goal.
We will prove w' < x.
An exact proof term for the current goal is Hx3 w' Hw'.
We will prove x < z0.
An exact proof term for the current goal is Hz0c.
We will prove z'Rx, z0 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLtLe_or z0 z' Hz0a (HLRx2 z' Hz') to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: z' z0.
We will prove False.
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz'.
We will prove z' z0.
An exact proof term for the current goal is H.
Apply L8a to the current goal.
Assume Hxz0: SNoLev x SNoLev z0.
Assume _.
Apply In_irref (SNoLev z0) to the current goal.
We will prove SNoLev z0 SNoLev z0.
Apply Hxz0 to the current goal.
An exact proof term for the current goal is Hz0b.
We prove the intermediate claim L9: w1'Ly, w1 w1'.
Apply dneg to the current goal.
Assume HC.
We prove the intermediate claim L9a: SNoLev y SNoLev w1 SNoEq_ (SNoLev y) y w1.
Apply Hy5 w1 Hw1a to the current goal.
We will prove w'Ly, w' < w1.
Let w' be given.
Assume Hw'.
Apply SNoLtLe_or w' w1 (HLRy1 w' Hw') Hw1a to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: w1 w'.
We will prove False.
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw'.
We will prove w1 w'.
An exact proof term for the current goal is H.
We will prove z'Ry, w1 < z'.
Let z' be given.
Assume Hz'.
Apply SNoLt_tra w1 y z' Hw1a Hy1 (HLRy2 z' Hz') to the current goal.
We will prove w1 < y.
An exact proof term for the current goal is Hw1c.
We will prove y < z'.
An exact proof term for the current goal is Hy4 z' Hz'.
Apply L9a to the current goal.
Assume Hyw1: SNoLev y SNoLev w1.
Assume _.
Apply In_irref (SNoLev w1) to the current goal.
We will prove SNoLev w1 SNoLev w1.
Apply Hyw1 to the current goal.
An exact proof term for the current goal is Hw1b.
Apply L8 to the current goal.
Let z0' be given.
Assume H.
Apply H to the current goal.
Assume Hz0'1: z0' Rx.
Assume Hz0'2: z0' z0.
Apply L9 to the current goal.
Let w1' be given.
Assume H.
Apply H to the current goal.
Assume Hw1'1: w1' Ly.
Assume Hw1'2: w1 w1'.
We will prove v < z0 * y + x * w1 + - z0 * w1.
Apply SNoLtLe_tra v (z0' * y + x * w1' + - z0' * w1') (z0 * y + x * w1 + - z0 * w1) Hv1 (SNo_add_SNo_3 (z0' * y) (x * w1') (- z0' * w1') (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_minus_SNo (z0' * w1') (SNo_mul_SNo z0' w1' (HLRx2 z0' Hz0'1) (HLRy1 w1' Hw1'1)))) (SNo_add_SNo_3 (z0 * y) (x * w1) (- z0 * w1) (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_minus_SNo (z0 * w1) (SNo_mul_SNo z0 w1 Hz0a Hw1a))) to the current goal.
We will prove v < z0' * y + x * w1' + - z0' * w1'.
Apply Hv4 to the current goal.
We will prove z0' * y + x * w1' + - z0' * w1' LxRy' RxLy'.
Apply binunionI2 to the current goal.
Apply RxLy'I to the current goal.
An exact proof term for the current goal is Hz0'1.
An exact proof term for the current goal is Hw1'1.
We will prove z0' * y + x * w1' + - z0' * w1' z0 * y + x * w1 + - z0 * w1.
Apply add_SNo_minus_Le12b3 (z0' * y) (x * w1') (z0' * w1') (z0 * y) (x * w1) (z0 * w1) (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo z0' w1' (HLRx2 z0' Hz0'1) (HLRy1 w1' Hw1'1)) (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo z0 w1 Hz0a Hw1a) to the current goal.
We will prove z0' * y + x * w1' + z0 * w1 z0 * y + x * w1 + z0' * w1'.
Apply SNoLe_tra (z0' * y + x * w1' + z0 * w1) (z0' * y + x * w1 + z0 * w1') (z0 * y + x * w1 + z0' * w1') (SNo_add_SNo_3 (z0' * y) (x * w1') (z0 * w1) (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo z0 w1 Hz0a Hw1a)) (SNo_add_SNo_3 (z0' * y) (x * w1) (z0 * w1') (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo z0 w1' Hz0a (HLRy1 w1' Hw1'1))) (SNo_add_SNo_3 (z0 * y) (x * w1) (z0' * w1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo z0' w1' (HLRx2 z0' Hz0'1) (HLRy1 w1' Hw1'1))) to the current goal.
We will prove z0' * y + x * w1' + z0 * w1 z0' * y + x * w1 + z0 * w1'.
Apply add_SNo_Le2 (z0' * y) (x * w1' + z0 * w1) (x * w1 + z0 * w1') (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_add_SNo (x * w1') (z0 * w1) (SNo_mul_SNo x w1' Hx1 (HLRy1 w1' Hw1'1)) (SNo_mul_SNo z0 w1 Hz0a Hw1a)) (SNo_add_SNo (x * w1) (z0 * w1') (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo z0 w1' Hz0a (HLRy1 w1' Hw1'1))) to the current goal.
We will prove x * w1' + z0 * w1 x * w1 + z0 * w1'.
rewrite the current goal using add_SNo_com (x * w1) (z0 * w1') (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo z0 w1' Hz0a (HLRy1 w1' Hw1'1)) (from left to right).
Apply mul_SNo_Le z0 w1' x w1 Hz0a (HLRy1 w1' Hw1'1) Hx1 Hw1a to the current goal.
We will prove x z0.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hz0c.
We will prove w1 w1'.
An exact proof term for the current goal is Hw1'2.
We will prove z0' * y + x * w1 + z0 * w1' z0 * y + x * w1 + z0' * w1'.
rewrite the current goal using add_SNo_com_3_0_1 (z0' * y) (x * w1) (z0 * w1') (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo z0 w1' Hz0a (HLRy1 w1' Hw1'1)) (from left to right).
rewrite the current goal using add_SNo_com_3_0_1 (z0 * y) (x * w1) (z0' * w1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo z0' w1' (HLRx2 z0' Hz0'1) (HLRy1 w1' Hw1'1)) (from left to right).
We will prove x * w1 + z0' * y + z0 * w1' x * w1 + z0 * y + z0' * w1'.
Apply add_SNo_Le2 (x * w1) (z0' * y + z0 * w1') (z0 * y + z0' * w1') (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_add_SNo (z0' * y) (z0 * w1') (SNo_mul_SNo z0' y (HLRx2 z0' Hz0'1) Hy1) (SNo_mul_SNo z0 w1' Hz0a (HLRy1 w1' Hw1'1))) (SNo_add_SNo (z0 * y) (z0' * w1') (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo z0' w1' (HLRx2 z0' Hz0'1) (HLRy1 w1' Hw1'1))) to the current goal.
We will prove z0' * y + z0 * w1' z0 * y + z0' * w1'.
Apply mul_SNo_Le z0 y z0' w1' Hz0a Hy1 (HLRx2 z0' Hz0'1) (HLRy1 w1' Hw1'1) to the current goal.
We will prove z0' z0.
An exact proof term for the current goal is Hz0'2.
We will prove w1' y.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is Hy3 w1' Hw1'1.
rewrite the current goal using HLR'eq (from right to left).
We will prove wLxLy' RxRy', w < x * y.
Let w be given.
Apply binunionE' to the current goal.
Assume Hw: w LxLy'.
Apply LxLy'E w Hw to the current goal.
Let w' be given.
Assume Hw': w' Lx.
Let w'' be given.
Assume Hw'': w'' Ly.
Assume Hw'1: SNo w'.
Assume Hw''1: SNo w''.
Assume Hw'2: w' < x.
Assume Hw''2: w'' < y.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
We will prove w' * y + x * w'' + - w' * w'' < x * y.
Apply add_SNo_minus_Lt1b3 (w' * y) (x * w'') (w' * w'') (x * y) (SNo_mul_SNo w' y Hw'1 Hy1) (SNo_mul_SNo x w'' Hx1 Hw''1) (SNo_mul_SNo w' w'' Hw'1 Hw''1) (SNo_mul_SNo x y Hx1 Hy1) to the current goal.
We will prove w' * y + x * w'' < x * y + w' * w''.
An exact proof term for the current goal is mul_SNo_Lt x y w' w'' Hx1 Hy1 Hw'1 Hw''1 Hw'2 Hw''2.
Assume Hw: w RxRy'.
Apply RxRy'E w Hw to the current goal.
Let z be given.
Assume Hz: z Rx.
Let z' be given.
Assume Hz': z' Ry.
Assume Hz1: SNo z.
Assume Hz'1: SNo z'.
Assume Hz2: x < z.
Assume Hz'2: y < z'.
Assume Hwe.
rewrite the current goal using Hwe (from left to right).
We will prove z * y + x * z' + - z * z' < x * y.
Apply add_SNo_minus_Lt1b3 (z * y) (x * z') (z * z') (x * y) (SNo_mul_SNo z y Hz1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (SNo_mul_SNo z z' Hz1 Hz'1) (SNo_mul_SNo x y Hx1 Hy1) to the current goal.
We will prove z * y + x * z' < x * y + z * z'.
rewrite the current goal using add_SNo_com (z * y) (x * z') (SNo_mul_SNo z y Hz1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (from left to right).
rewrite the current goal using add_SNo_com (x * y) (z * z') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo z z' Hz1 Hz'1) (from left to right).
An exact proof term for the current goal is mul_SNo_Lt z z' x y Hz1 Hz'1 Hx1 Hy1 Hz2 Hz'2.
rewrite the current goal using HLR'eq (from right to left).
We will prove zLxRy' RxLy', x * y < z.
Let z be given.
Apply binunionE' to the current goal.
Assume Hz: z LxRy'.
Apply LxRy'E z Hz to the current goal.
Let w be given.
Assume Hw: w Lx.
Let z' be given.
Assume Hz': z' Ry.
Assume Hw1: SNo w.
Assume Hz'1: SNo z'.
Assume Hw2: w < x.
Assume Hz'2: y < z'.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove x * y < w * y + x * z' + - w * z'.
Apply add_SNo_minus_Lt2b3 (w * y) (x * z') (w * z') (x * y) (SNo_mul_SNo w y Hw1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (SNo_mul_SNo w z' Hw1 Hz'1) (SNo_mul_SNo x y Hx1 Hy1) to the current goal.
We will prove x * y + w * z' < w * y + x * z'.
rewrite the current goal using add_SNo_com (x * y) (w * z') (SNo_mul_SNo x y Hx1 Hy1) (SNo_mul_SNo w z' Hw1 Hz'1) (from left to right).
rewrite the current goal using add_SNo_com (w * y) (x * z') (SNo_mul_SNo w y Hw1 Hy1) (SNo_mul_SNo x z' Hx1 Hz'1) (from left to right).
We will prove w * z' + x * y < x * z' + w * y.
An exact proof term for the current goal is mul_SNo_Lt x z' w y Hx1 Hz'1 Hw1 Hy1 Hw2 Hz'2.
Assume Hz: z RxLy'.
Apply RxLy'E z Hz to the current goal.
Let z' be given.
Assume Hz': z' Rx.
Let w be given.
Assume Hw: w Ly.
Assume Hz'1: SNo z'.
Assume Hw1: SNo w.
Assume Hz'2: x < z'.
Assume Hw2: w < y.
Assume Hze.
rewrite the current goal using Hze (from left to right).
We will prove x * y < z' * y + x * w + - z' * w.
Apply add_SNo_minus_Lt2b3 (z' * y) (x * w) (z' * w) (x * y) (SNo_mul_SNo z' y Hz'1 Hy1) (SNo_mul_SNo x w Hx1 Hw1) (SNo_mul_SNo z' w Hz'1 Hw1) (SNo_mul_SNo x y Hx1 Hy1) to the current goal.
We will prove x * y + z' * w < z' * y + x * w.
An exact proof term for the current goal is mul_SNo_Lt z' y x w Hz'1 Hy1 Hx1 Hw1 Hz'2 Hw2.
Apply and3I to the current goal.
An exact proof term for the current goal is L1.
An exact proof term for the current goal is Lxyeq.
Let q be given.
Assume Hq.
Apply Hq LxLy' RxRy' LxRy' RxLy' to the current goal.
An exact proof term for the current goal is LxLy'E.
An exact proof term for the current goal is LxLy'I.
An exact proof term for the current goal is RxRy'E.
An exact proof term for the current goal is RxRy'I.
An exact proof term for the current goal is LxRy'E.
An exact proof term for the current goal is LxRy'I.
An exact proof term for the current goal is RxLy'E.
An exact proof term for the current goal is RxLy'I.
An exact proof term for the current goal is L1.
An exact proof term for the current goal is Lxyeq.
Theorem. (mul_SNoCutP_gen) The following is provable:
∀Lx Rx Ly Ry x y, SNoCutP Lx RxSNoCutP Ly Ryx = SNoCut Lx Rxy = SNoCut Ly RySNoCutP ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ly} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ry}) ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ry} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ly})
Proof:
Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye.
Apply mul_SNoCutP_lem Lx Rx Ly Ry x y HLRx HLRy Hxe Hye to the current goal.
Assume H _.
Apply H to the current goal.
Assume H _.
An exact proof term for the current goal is H.
Theorem. (mul_SNoCut_eq) The following is provable:
∀Lx Rx Ly Ry x y, SNoCutP Lx RxSNoCutP Ly Ryx = SNoCut Lx Rxy = SNoCut Ly Ryx * y = SNoCut ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ly} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ry}) ({w 0 * y + x * w 1 + - w 0 * w 1|wLx Ry} {z 0 * y + x * z 1 + - z 0 * z 1|zRx Ly})
Proof:
Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye.
Apply mul_SNoCutP_lem Lx Rx Ly Ry x y HLRx HLRy Hxe Hye to the current goal.
Assume H _.
Apply H to the current goal.
Assume _ H.
An exact proof term for the current goal is H.
Theorem. (mul_SNoCut_abs) The following is provable:
∀Lx Rx Ly Ry x y, SNoCutP Lx RxSNoCutP Ly Ryx = SNoCut Lx Rxy = SNoCut Ly Ry∀q : prop, (∀LxLy' RxRy' LxRy' RxLy', (uLxLy', ∀p : prop, (wLx, w'Ly, SNo wSNo w'w < xw' < yu = w * y + x * w' + - w * w'p)p)(wLx, w'Ly, w * y + x * w' + - w * w' LxLy')(uRxRy', ∀p : prop, (zRx, z'Ry, SNo zSNo z'x < zy < z'u = z * y + x * z' + - z * z'p)p)(zRx, z'Ry, z * y + x * z' + - z * z' RxRy')(uLxRy', ∀p : prop, (wLx, zRy, SNo wSNo zw < xy < zu = w * y + x * z + - w * zp)p)(wLx, zRy, w * y + x * z + - w * z LxRy')(uRxLy', ∀p : prop, (zRx, wLy, SNo zSNo wx < zw < yu = z * y + x * w + - z * wp)p)(zRx, wLy, z * y + x * w + - z * w RxLy')SNoCutP (LxLy' RxRy') (LxRy' RxLy')x * y = SNoCut (LxLy' RxRy') (LxRy' RxLy')q)q
Proof:
Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye.
Apply mul_SNoCutP_lem Lx Rx Ly Ry x y HLRx HLRy Hxe Hye to the current goal.
Assume _ H.
An exact proof term for the current goal is H.
Theorem. (mul_SNo_SNoCut_SNoL_interpolate) The following is provable:
∀Lx Rx Ly Ry, SNoCutP Lx RxSNoCutP Ly Ry∀x y, x = SNoCut Lx Rxy = SNoCut Ly RyuSNoL (x * y), (vLx, wLy, u + v * w v * y + x * w) (vRx, wRy, u + v * w v * y + x * w)
Proof:
Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy.
Let x and y be given.
Assume Hx Hy.
Set P1 to be the term λu ⇒ vLx, wLy, u + v * w v * y + x * w of type setprop.
Set P2 to be the term λu ⇒ vRx, wRy, u + v * w v * y + x * w of type setprop.
Set P to be the term λu ⇒ P1 u P2 u of type setprop.
Apply HLRx to the current goal.
Assume H.
Apply H to the current goal.
Assume HLx HRx HLRx'.
Apply HLRy to the current goal.
Assume H.
Apply H to the current goal.
Assume HLy HRy HLRy'.
Apply SNoCutP_SNoCut_impred Lx Rx HLRx to the current goal.
rewrite the current goal using Hx (from right to left).
Assume Hx1: SNo x.
Assume Hx2: SNoLev x ordsucc ((xLxordsucc (SNoLev x)) (yRxordsucc (SNoLev y))).
Assume Hx3: wLx, w < x.
Assume Hx4: zRx, x < z.
Assume Hx5: ∀u, SNo u(wLx, w < u)(zRx, u < z)SNoLev x SNoLev u SNoEq_ (SNoLev x) x u.
Apply SNoCutP_SNoCut_impred Ly Ry HLRy to the current goal.
rewrite the current goal using Hy (from right to left).
Assume Hy1: SNo y.
Assume Hy2: SNoLev y ordsucc ((yLyordsucc (SNoLev y)) (yRyordsucc (SNoLev y))).
Assume Hy3: wLy, w < y.
Assume Hy4: zRy, y < z.
Assume Hy5: ∀u, SNo u(wLy, w < u)(zRy, u < z)SNoLev y SNoLev u SNoEq_ (SNoLev y) y u.
We prove the intermediate claim Lxy: SNo (x * y).
An exact proof term for the current goal is SNo_mul_SNo x y Hx1 Hy1.
We prove the intermediate claim LI: ∀u, SNo uSNoLev u SNoLev (x * y)u < x * yP u.
Apply SNoLev_ind to the current goal.
Let u be given.
Assume Hu1: SNo u.
Assume IH: zSNoS_ (SNoLev u), SNoLev z SNoLev (x * y)z < x * yP z.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: u < x * y.
Apply dneg to the current goal.
Assume HNC: ¬ P u.
We prove the intermediate claim L1: vLx, wLy, v * y + x * w < u + v * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HLx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HLy w Hw.
Apply SNoLtLe_or (v * y + x * w) (u + v * w) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Lv1 Lw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
We prove the intermediate claim L2: vRx, wRy, v * y + x * w < u + v * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HRx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HRy w Hw.
Apply SNoLtLe_or (v * y + x * w) (u + v * w) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Lv1 Lw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
Apply SNoLt_irref u to the current goal.
Apply SNoLtLe_tra u (x * y) u Hu1 Lxy Hu1 Hu3 to the current goal.
We will prove x * y u.
Apply mul_SNoCut_abs Lx Rx Ly Ry x y HLRx HLRy Hx Hy to the current goal.
Let LxLy', RxRy', LxRy' and RxLy' be given.
Assume LxLy'E LxLy'I RxRy'E RxRy'I LxRy'E LxRy'I RxLy'E RxLy'I.
Assume HSC: SNoCutP (LxLy' RxRy') (LxRy' RxLy').
Assume HE: x * y = SNoCut (LxLy' RxRy') (LxRy' RxLy').
rewrite the current goal using HE (from left to right).
We will prove SNoCut (LxLy' RxRy') (LxRy' RxLy') u.
rewrite the current goal using SNo_eta u Hu1 (from left to right).
We will prove SNoCut (LxLy' RxRy') (LxRy' RxLy') SNoCut (SNoL u) (SNoR u).
Apply SNoCut_Le (LxLy' RxRy') (LxRy' RxLy') (SNoL u) (SNoR u) HSC (SNoCutP_SNoL_SNoR u Hu1) to the current goal.
We will prove zLxLy' RxRy', z < SNoCut (SNoL u) (SNoR u).
Let z be given.
rewrite the current goal using SNo_eta u Hu1 (from right to left).
Apply binunionE' to the current goal.
Assume Hz: z LxLy'.
Apply LxLy'E z Hz to the current goal.
Let w0 be given.
Assume Hw0.
Let w1 be given.
Assume Hw1.
Assume Hw0a Hw1a Hw0x Hw1y Hze.
rewrite the current goal using Hze (from left to right).
We will prove w0 * y + x * w1 + - w0 * w1 < u.
Apply add_SNo_minus_Lt1b3 (w0 * y) (x * w1) (w0 * w1) u (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo w0 w1 Hw0a Hw1a) Hu1 to the current goal.
We will prove w0 * y + x * w1 < u + w0 * w1.
Apply SNoLtLe_or (w0 * y + x * w1) (u + w0 * w1) (SNo_add_SNo (w0 * y) (x * w1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a)) (SNo_add_SNo u (w0 * w1) Hu1 (SNo_mul_SNo w0 w1 Hw0a Hw1a)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: u + w0 * w1 w0 * y + x * w1.
Apply HNC to the current goal.
We will prove P1 u P2 u.
Apply orIL to the current goal.
We will prove vLx, wLy, u + v * w v * y + x * w.
We use w0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw0.
We use w1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw1.
An exact proof term for the current goal is H.
Assume Hz: z RxRy'.
Apply RxRy'E z Hz to the current goal.
Let z0 be given.
Assume Hz0.
Let z1 be given.
Assume Hz1.
Assume Hz0a Hz1a Hz0x Hz1y Hze.
rewrite the current goal using Hze (from left to right).
We will prove z0 * y + x * z1 + - z0 * z1 < u.
Apply add_SNo_minus_Lt1b3 (z0 * y) (x * z1) (z0 * z1) u (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo z0 z1 Hz0a Hz1a) Hu1 to the current goal.
We will prove z0 * y + x * z1 < u + z0 * z1.
Apply SNoLtLe_or (z0 * y + x * z1) (u + z0 * z1) (SNo_add_SNo (z0 * y) (x * z1) (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a)) (SNo_add_SNo u (z0 * z1) Hu1 (SNo_mul_SNo z0 z1 Hz0a Hz1a)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: u + z0 * z1 z0 * y + x * z1.
Apply HNC to the current goal.
We will prove P1 u P2 u.
Apply orIR to the current goal.
We will prove vRx, wRy, u + v * w v * y + x * w.
We use z0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz0.
We use z1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz1.
An exact proof term for the current goal is H.
We will prove zSNoR u, SNoCut (LxLy' RxRy') (LxRy' RxLy') < z.
Let z be given.
Assume Hz: z SNoR u.
rewrite the current goal using HE (from right to left).
We will prove x * y < z.
Apply SNoR_E u Hu1 z Hz to the current goal.
Assume Hz1: SNo z.
Assume Hz2: SNoLev z SNoLev u.
Assume Hz3: u < z.
Apply SNoLt_trichotomy_or_impred z (x * y) Hz1 Lxy to the current goal.
Assume H1: z < x * y.
We prove the intermediate claim LPz: P z.
Apply IH z to the current goal.
We will prove z SNoS_ (SNoLev u).
Apply SNoS_I2 to the current goal.
An exact proof term for the current goal is Hz1.
An exact proof term for the current goal is Hu1.
An exact proof term for the current goal is Hz2.
We will prove SNoLev z SNoLev (x * y).
An exact proof term for the current goal is ordinal_TransSet (SNoLev (x * y)) (SNoLev_ordinal (x * y) Lxy) (SNoLev u) Hu2 (SNoLev z) Hz2.
We will prove z < x * y.
An exact proof term for the current goal is H1.
Apply LPz to the current goal.
Assume H2: P1 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v Lx.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w Ly.
Assume Hvw: z + v * w v * y + x * w.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HLx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HLy w Hw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Lv1 Lw1.
We prove the intermediate claim L3: z + v * w < u + v * w.
Apply SNoLeLt_tra (z + v * w) (v * y + x * w) (u + v * w) (SNo_add_SNo z (v * w) Hz1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) (SNo_add_SNo u (v * w) Hu1 Lvw) Hvw to the current goal.
We will prove v * y + x * w < u + v * w.
An exact proof term for the current goal is L1 v Hv w Hw.
We prove the intermediate claim L4: z < u.
An exact proof term for the current goal is add_SNo_Lt1_cancel z (v * w) u Hz1 Lvw Hu1 L3.
We will prove False.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L4.
Assume H2: P2 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v Rx.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w Ry.
Assume Hvw: z + v * w v * y + x * w.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HRx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HRy w Hw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Lv1 Lw1.
We prove the intermediate claim L5: z + v * w < u + v * w.
Apply SNoLeLt_tra (z + v * w) (v * y + x * w) (u + v * w) (SNo_add_SNo z (v * w) Hz1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) (SNo_add_SNo u (v * w) Hu1 Lvw) Hvw to the current goal.
We will prove v * y + x * w < u + v * w.
An exact proof term for the current goal is L2 v Hv w Hw.
We prove the intermediate claim L6: z < u.
An exact proof term for the current goal is add_SNo_Lt1_cancel z (v * w) u Hz1 Lvw Hu1 L5.
We will prove False.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L6.
Assume H1: z = x * y.
Apply In_no2cycle (SNoLev u) (SNoLev (x * y)) Hu2 to the current goal.
We will prove SNoLev (x * y) SNoLev u.
rewrite the current goal using H1 (from right to left).
An exact proof term for the current goal is Hz2.
Assume H1: x * y < z.
An exact proof term for the current goal is H1.
Let u be given.
Assume Hu: u SNoL (x * y).
Apply SNoL_E (x * y) Lxy u Hu to the current goal.
Assume Hu1: SNo u.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: u < x * y.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3.
Theorem. (mul_SNo_SNoCut_SNoL_interpolate_impred) The following is provable:
∀Lx Rx Ly Ry, SNoCutP Lx RxSNoCutP Ly Ry∀x y, x = SNoCut Lx Rxy = SNoCut Ly RyuSNoL (x * y), ∀p : prop, (vLx, wLy, u + v * w v * y + x * wp)(vRx, wRy, u + v * w v * y + x * wp)p
Proof:
Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy.
Let x and y be given.
Assume Hx Hy.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp1 Hp2.
Apply mul_SNo_SNoCut_SNoL_interpolate Lx Rx Ly Ry HLRx HLRy x y Hx Hy u Hu to the current goal.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw.
Theorem. (mul_SNo_SNoCut_SNoR_interpolate) The following is provable:
∀Lx Rx Ly Ry, SNoCutP Lx RxSNoCutP Ly Ry∀x y, x = SNoCut Lx Rxy = SNoCut Ly RyuSNoR (x * y), (vLx, wRy, v * y + x * w u + v * w) (vRx, wLy, v * y + x * w u + v * w)
Proof:
Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy.
Let x and y be given.
Assume Hx Hy.
Set P1 to be the term λu ⇒ vLx, wRy, v * y + x * w u + v * w of type setprop.
Set P2 to be the term λu ⇒ vRx, wLy, v * y + x * w u + v * w of type setprop.
Set P to be the term λu ⇒ P1 u P2 u of type setprop.
Apply HLRx to the current goal.
Assume H.
Apply H to the current goal.
Assume HLx HRx HLRx'.
Apply HLRy to the current goal.
Assume H.
Apply H to the current goal.
Assume HLy HRy HLRy'.
Apply SNoCutP_SNoCut_impred Lx Rx HLRx to the current goal.
rewrite the current goal using Hx (from right to left).
Assume Hx1: SNo x.
Assume Hx2: SNoLev x ordsucc ((xLxordsucc (SNoLev x)) (yRxordsucc (SNoLev y))).
Assume Hx3: wLx, w < x.
Assume Hx4: zRx, x < z.
Assume Hx5: ∀u, SNo u(wLx, w < u)(zRx, u < z)SNoLev x SNoLev u SNoEq_ (SNoLev x) x u.
Apply SNoCutP_SNoCut_impred Ly Ry HLRy to the current goal.
rewrite the current goal using Hy (from right to left).
Assume Hy1: SNo y.
Assume Hy2: SNoLev y ordsucc ((yLyordsucc (SNoLev y)) (yRyordsucc (SNoLev y))).
Assume Hy3: wLy, w < y.
Assume Hy4: zRy, y < z.
Assume Hy5: ∀u, SNo u(wLy, w < u)(zRy, u < z)SNoLev y SNoLev u SNoEq_ (SNoLev y) y u.
We prove the intermediate claim Lxy: SNo (x * y).
An exact proof term for the current goal is SNo_mul_SNo x y Hx1 Hy1.
We prove the intermediate claim LI: ∀u, SNo uSNoLev u SNoLev (x * y)x * y < uP u.
Apply SNoLev_ind to the current goal.
Let u be given.
Assume Hu1: SNo u.
Assume IH: zSNoS_ (SNoLev u), SNoLev z SNoLev (x * y)x * y < zP z.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: x * y < u.
Apply dneg to the current goal.
Assume HNC: ¬ P u.
We prove the intermediate claim L1: vLx, wRy, u + v * w < v * y + x * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HLx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HRy w Hw.
Apply SNoLtLe_or (u + v * w) (v * y + x * w) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Lv1 Lw1)) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
We prove the intermediate claim L2: vRx, wLy, u + v * w < v * y + x * w.
Let v be given.
Assume Hv.
Let w be given.
Assume Hw.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HRx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HLy w Hw.
Apply SNoLtLe_or (u + v * w) (v * y + x * w) (SNo_add_SNo u (v * w) Hu1 (SNo_mul_SNo v w Lv1 Lw1)) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H.
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv.
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw.
An exact proof term for the current goal is H.
Apply SNoLt_irref (x * y) to the current goal.
Apply SNoLtLe_tra (x * y) u (x * y) Lxy Hu1 Lxy Hu3 to the current goal.
We will prove u x * y.
Apply mul_SNoCut_abs Lx Rx Ly Ry x y HLRx HLRy Hx Hy to the current goal.
Let LxLy', RxRy', LxRy' and RxLy' be given.
Assume LxLy'E LxLy'I RxRy'E RxRy'I LxRy'E LxRy'I RxLy'E RxLy'I.
Assume HSC: SNoCutP (LxLy' RxRy') (LxRy' RxLy').
Assume HE: x * y = SNoCut (LxLy' RxRy') (LxRy' RxLy').
rewrite the current goal using HE (from left to right).
We will prove u SNoCut (LxLy' RxRy') (LxRy' RxLy').
rewrite the current goal using SNo_eta u Hu1 (from left to right).
We will prove SNoCut (SNoL u) (SNoR u) SNoCut (LxLy' RxRy') (LxRy' RxLy').
Apply SNoCut_Le (SNoL u) (SNoR u) (LxLy' RxRy') (LxRy' RxLy') (SNoCutP_SNoL_SNoR u Hu1) HSC to the current goal.
We will prove zSNoL u, z < SNoCut (LxLy' RxRy') (LxRy' RxLy').
Let z be given.
Assume Hz: z SNoL u.
rewrite the current goal using HE (from right to left).
We will prove z < x * y.
Apply SNoL_E u Hu1 z Hz to the current goal.
Assume Hz1: SNo z.
Assume Hz2: SNoLev z SNoLev u.
Assume Hz3: z < u.
Apply SNoLt_trichotomy_or_impred (x * y) z Lxy Hz1 to the current goal.
Assume H1: x * y < z.
We prove the intermediate claim LPz: P z.
Apply IH z to the current goal.
We will prove z SNoS_ (SNoLev u).
Apply SNoS_I2 to the current goal.
An exact proof term for the current goal is Hz1.
An exact proof term for the current goal is Hu1.
An exact proof term for the current goal is Hz2.
We will prove SNoLev z SNoLev (x * y).
An exact proof term for the current goal is ordinal_TransSet (SNoLev (x * y)) (SNoLev_ordinal (x * y) Lxy) (SNoLev u) Hu2 (SNoLev z) Hz2.
We will prove x * y < z.
An exact proof term for the current goal is H1.
Apply LPz to the current goal.
Assume H2: P1 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v Lx.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w Ry.
Assume Hvw: v * y + x * w z + v * w.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HLx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HRy w Hw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Lv1 Lw1.
We prove the intermediate claim L3: u + v * w < z + v * w.
Apply SNoLtLe_tra (u + v * w) (v * y + x * w) (z + v * w) (SNo_add_SNo u (v * w) Hu1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) (SNo_add_SNo z (v * w) Hz1 Lvw) to the current goal.
We will prove u + v * w < v * y + x * w.
An exact proof term for the current goal is L1 v Hv w Hw.
An exact proof term for the current goal is Hvw.
We prove the intermediate claim L4: u < z.
An exact proof term for the current goal is add_SNo_Lt1_cancel u (v * w) z Hu1 Lvw Hz1 L3.
We will prove False.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 L4 Hz3.
Assume H2: P2 z.
Apply H2 to the current goal.
Let v be given.
Assume H2.
Apply H2 to the current goal.
Assume Hv: v Rx.
Assume H2.
Apply H2 to the current goal.
Let w be given.
Assume H2.
Apply H2 to the current goal.
Assume Hw: w Ly.
Assume Hvw: v * y + x * w z + v * w.
We prove the intermediate claim Lv1: SNo v.
An exact proof term for the current goal is HRx v Hv.
We prove the intermediate claim Lw1: SNo w.
An exact proof term for the current goal is HLy w Hw.
We prove the intermediate claim Lvw: SNo (v * w).
An exact proof term for the current goal is SNo_mul_SNo v w Lv1 Lw1.
We prove the intermediate claim L5: u + v * w < z + v * w.
Apply SNoLtLe_tra (u + v * w) (v * y + x * w) (z + v * w) (SNo_add_SNo u (v * w) Hu1 Lvw) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Lv1 Hy1) (SNo_mul_SNo x w Hx1 Lw1)) (SNo_add_SNo z (v * w) Hz1 Lvw) to the current goal.
We will prove u + v * w < v * y + x * w.
An exact proof term for the current goal is L2 v Hv w Hw.
An exact proof term for the current goal is Hvw.
We prove the intermediate claim L6: u < z.
An exact proof term for the current goal is add_SNo_Lt1_cancel u (v * w) z Hu1 Lvw Hz1 L5.
We will prove False.
Apply SNoLt_irref u to the current goal.
An exact proof term for the current goal is SNoLt_tra u z u Hu1 Hz1 Hu1 L6 Hz3.
Assume H1: x * y = z.
Apply In_no2cycle (SNoLev u) (SNoLev (x * y)) Hu2 to the current goal.
We will prove SNoLev (x * y) SNoLev u.
rewrite the current goal using H1 (from left to right).
An exact proof term for the current goal is Hz2.
Assume H1: z < x * y.
An exact proof term for the current goal is H1.
We will prove zLxRy' RxLy', SNoCut (SNoL u) (SNoR u) < z.
Let z be given.
rewrite the current goal using SNo_eta u Hu1 (from right to left).
Apply binunionE' to the current goal.
Assume Hz: z LxRy'.
Apply LxRy'E z Hz to the current goal.
Let w0 be given.
Assume Hw0.
Let w1 be given.
Assume Hw1.
Assume Hw0a Hw1a Hw0x Hw1y Hze.
rewrite the current goal using Hze (from left to right).
We will prove u < w0 * y + x * w1 + - w0 * w1.
Apply add_SNo_minus_Lt2b3 (w0 * y) (x * w1) (w0 * w1) u (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a) (SNo_mul_SNo w0 w1 Hw0a Hw1a) Hu1 to the current goal.
We will prove u + w0 * w1 < w0 * y + x * w1.
Apply SNoLtLe_or (u + w0 * w1) (w0 * y + x * w1) (SNo_add_SNo u (w0 * w1) Hu1 (SNo_mul_SNo w0 w1 Hw0a Hw1a)) (SNo_add_SNo (w0 * y) (x * w1) (SNo_mul_SNo w0 y Hw0a Hy1) (SNo_mul_SNo x w1 Hx1 Hw1a)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: w0 * y + x * w1 u + w0 * w1.
Apply HNC to the current goal.
We will prove P1 u P2 u.
Apply orIL to the current goal.
We will prove vLx, wRy, v * y + x * w u + v * w.
We use w0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw0.
We use w1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw1.
An exact proof term for the current goal is H.
Assume Hz: z RxLy'.
Apply RxLy'E z Hz to the current goal.
Let z0 be given.
Assume Hz0.
Let z1 be given.
Assume Hz1.
Assume Hz0a Hz1a Hz0x Hz1y Hze.
rewrite the current goal using Hze (from left to right).
We will prove u < z0 * y + x * z1 + - z0 * z1.
Apply add_SNo_minus_Lt2b3 (z0 * y) (x * z1) (z0 * z1) u (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a) (SNo_mul_SNo z0 z1 Hz0a Hz1a) Hu1 to the current goal.
We will prove u + z0 * z1 < z0 * y + x * z1.
Apply SNoLtLe_or (u + z0 * z1) (z0 * y + x * z1) (SNo_add_SNo u (z0 * z1) Hu1 (SNo_mul_SNo z0 z1 Hz0a Hz1a)) (SNo_add_SNo (z0 * y) (x * z1) (SNo_mul_SNo z0 y Hz0a Hy1) (SNo_mul_SNo x z1 Hx1 Hz1a)) to the current goal.
Assume H.
An exact proof term for the current goal is H.
Assume H: z0 * y + x * z1 u + z0 * z1.
Apply HNC to the current goal.
We will prove P1 u P2 u.
Apply orIR to the current goal.
We will prove vRx, wLy, v * y + x * w u + v * w.
We use z0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz0.
We use z1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz1.
An exact proof term for the current goal is H.
Let u be given.
Assume Hu: u SNoR (x * y).
Apply SNoR_E (x * y) Lxy u Hu to the current goal.
Assume Hu1: SNo u.
Assume Hu2: SNoLev u SNoLev (x * y).
Assume Hu3: x * y < u.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3.
Theorem. (mul_SNo_SNoCut_SNoR_interpolate_impred) The following is provable:
∀Lx Rx Ly Ry, SNoCutP Lx RxSNoCutP Ly Ry∀x y, x = SNoCut Lx Rxy = SNoCut Ly RyuSNoR (x * y), ∀p : prop, (vLx, wRy, v * y + x * w u + v * wp)(vRx, wLy, v * y + x * w u + v * wp)p
Proof:
Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy.
Let x and y be given.
Assume Hx Hy.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp1 Hp2.
Apply mul_SNo_SNoCut_SNoR_interpolate Lx Rx Ly Ry HLRx HLRy x y Hx Hy u Hu to the current goal.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw.
Assume H1.
Apply H1 to the current goal.
Let v be given.
Assume H1.
Apply H1 to the current goal.
Assume Hv.
Assume H1.
Apply H1 to the current goal.
Let w be given.
Assume H1.
Apply H1 to the current goal.
Assume Hw Hvw.
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw.
End of Section SurrealMul
Beginning of Section SurrealExp
Notation. We use - as a prefix operator with priority 358 corresponding to applying term minus_SNo.
Notation. We use + as an infix operator with priority 360 and which associates to the right corresponding to applying term add_SNo.
Notation. We use * as an infix operator with priority 355 and which associates to the right corresponding to applying term mul_SNo.
Definition. We define exp_SNo_nat to be λn m : setnat_primrec 1 (λ_ r ⇒ n * r) m of type setsetset.
Notation. We use ^ as an infix operator with priority 342 and which associates to the right corresponding to applying term exp_SNo_nat.
Theorem. (exp_SNo_nat_0) The following is provable:
∀x, SNo xx ^ 0 = 1
Proof:
Let x be given.
Assume Hx.
An exact proof term for the current goal is nat_primrec_0 1 (λ_ r ⇒ x * r).
Theorem. (exp_SNo_nat_S) The following is provable:
∀x, SNo x∀n, nat_p nx ^ (ordsucc n) = x * x ^ n
Proof:
Let x be given.
Assume Hx.
Let n be given.
Assume Hn.
An exact proof term for the current goal is nat_primrec_S 1 (λ_ r ⇒ x * r) n Hn.
Theorem. (exp_SNo_nat_1) The following is provable:
∀x, SNo xx ^ 1 = x
Proof:
Let x be given.
Assume Hx.
rewrite the current goal using exp_SNo_nat_S x Hx 0 nat_0 (from left to right).
We will prove x * x ^ 0 = x.
rewrite the current goal using exp_SNo_nat_0 x Hx (from left to right).
We will prove x * 1 = x.
An exact proof term for the current goal is mul_SNo_oneR x Hx.
Theorem. (exp_SNo_nat_2) The following is provable:
∀x, SNo xx ^ 2 = x * x
Proof:
Let x be given.
Assume Hx.
rewrite the current goal using exp_SNo_nat_S x Hx 1 nat_1 (from left to right).
We will prove x * x ^ 1 = x * x.
Use f_equal.
An exact proof term for the current goal is exp_SNo_nat_1 x Hx.
Theorem. (SNo_sqr_nonneg') The following is provable:
∀x, SNo x0 x ^ 2
Proof:
Let x be given.
Assume Hx.
rewrite the current goal using exp_SNo_nat_2 x Hx (from left to right).
An exact proof term for the current goal is SNo_sqr_nonneg x Hx.
Theorem. (SNo_zero_or_sqr_pos') The following is provable:
∀x, SNo xx = 0 0 < x ^ 2
Proof:
Let x be given.
Assume Hx.
rewrite the current goal using exp_SNo_nat_2 x Hx (from left to right).
An exact proof term for the current goal is SNo_zero_or_sqr_pos x Hx.
Theorem. (SNo_exp_SNo_nat) The following is provable:
∀x, SNo x∀n, nat_p nSNo (x ^ n)
Proof:
Let x be given.
Assume Hx.
Apply nat_ind to the current goal.
We will prove SNo (x ^ 0).
rewrite the current goal using exp_SNo_nat_0 x Hx (from left to right).
An exact proof term for the current goal is SNo_1.
Let n be given.
Assume Hn.
Assume IHn: SNo (x ^ n).
We will prove SNo (x ^ (ordsucc n)).
rewrite the current goal using exp_SNo_nat_S x Hx n Hn (from left to right).
We will prove SNo (x * x ^ n).
An exact proof term for the current goal is SNo_mul_SNo x (x ^ n) Hx IHn.
Theorem. (nat_exp_SNo_nat) The following is provable:
∀x, nat_p x∀n, nat_p nnat_p (x ^ n)
Proof:
Let x be given.
Assume Hx.
We prove the intermediate claim Lx: SNo x.
An exact proof term for the current goal is nat_p_SNo x Hx.
Apply nat_ind to the current goal.
We will prove nat_p (x ^ 0).
rewrite the current goal using exp_SNo_nat_0 x Lx (from left to right).
An exact proof term for the current goal is nat_1.
Let n be given.
Assume Hn.
Assume IHn: nat_p (x ^ n).
We will prove nat_p (x ^ (ordsucc n)).
rewrite the current goal using exp_SNo_nat_S x Lx n Hn (from left to right).
We will prove nat_p (x * x ^ n).
rewrite the current goal using mul_nat_mul_SNo x (nat_p_omega x Hx) (x ^ n) (nat_p_omega (x ^ n) IHn) (from right to left).
We will prove nat_p (mul_nat x (x ^ n)).
An exact proof term for the current goal is mul_nat_p x Hx (x ^ n) IHn.
Theorem. (eps_ordsucc_half_add) The following is provable:
∀n, nat_p neps_ (ordsucc n) + eps_ (ordsucc n) = eps_ n
Proof:
Apply nat_complete_ind to the current goal.
Let n be given.
Assume Hn.
Assume IH: mn, eps_ (ordsucc m) + eps_ (ordsucc m) = eps_ m.
We prove the intermediate claim Ln: n ω.
An exact proof term for the current goal is nat_p_omega n Hn.
rewrite the current goal using eps_SNoCut n Ln (from left to right).
Set x to be the term eps_ (ordsucc n).
We prove the intermediate claim Lx: SNo x.
An exact proof term for the current goal is SNo_eps_ (ordsucc n) (omega_ordsucc n Ln).
We prove the intermediate claim Lx2: 0 < x.
An exact proof term for the current goal is SNo_eps_pos (ordsucc n) (omega_ordsucc n Ln).
rewrite the current goal using add_SNo_eq x Lx x Lx (from left to right).
We will prove SNoCut ({w + x|wSNoL x} {x + w|wSNoL x}) ({z + x|zSNoR x} {x + z|zSNoR x}) = SNoCut {0} {eps_ m|mn}.
Apply SNoCut_ext ({w + x|wSNoL x} {x + w|wSNoL x}) ({z + x|zSNoR x} {x + z|zSNoR x}) {0} {eps_ m|mn} (add_SNo_SNoCutP x x Lx Lx) (eps_SNoCutP n Ln) to the current goal.
We prove the intermediate claim L1: wSNoL x, w + x < eps_ n x + w < eps_ n.
Let w be given.
Assume Hw: w SNoL x.
Apply SNoL_E x Lx w Hw to the current goal.
Assume Hw1: SNo w.
rewrite the current goal using SNoLev_eps_ (ordsucc n) (omega_ordsucc n Ln) (from left to right).
Assume Hw2: SNoLev w ordsucc (ordsucc n).
Assume Hw3: w < x.
We prove the intermediate claim Lw0: w 0.
Apply SNoLtLe_or 0 w SNo_0 Hw1 to the current goal.
Assume H1: 0 < w.
We will prove False.
Apply SNoLt_irref w to the current goal.
We will prove w < w.
Apply SNoLt_tra w x w Hw1 Lx Hw1 Hw3 to the current goal.
We will prove x < w.
Apply SNo_pos_eps_Lt (ordsucc n) (nat_ordsucc n Hn) w to the current goal.
We will prove w SNoS_ (ordsucc (ordsucc n)).
Apply SNoS_I (ordsucc (ordsucc n)) (nat_p_ordinal (ordsucc (ordsucc n)) (nat_ordsucc (ordsucc n) (nat_ordsucc n Hn))) w (SNoLev w) Hw2 to the current goal.
We will prove SNo_ (SNoLev w) w.
An exact proof term for the current goal is SNoLev_ w Hw1.
We will prove 0 < w.
An exact proof term for the current goal is H1.
Assume H1: w 0.
An exact proof term for the current goal is H1.
We prove the intermediate claim L1a: w + x < eps_ n.
Apply SNoLeLt_tra (w + x) (0 + x) (eps_ n) to the current goal.
An exact proof term for the current goal is SNo_add_SNo w x Hw1 Lx.
An exact proof term for the current goal is SNo_add_SNo 0 x SNo_0 Lx.
An exact proof term for the current goal is SNo_eps_ n Ln.
We will prove w + x 0 + x.
An exact proof term for the current goal is add_SNo_Le1 w x 0 Hw1 Lx SNo_0 Lw0.
We will prove 0 + x < eps_ n.
rewrite the current goal using add_SNo_0L x Lx (from left to right).
We will prove eps_ (ordsucc n) < eps_ n.
An exact proof term for the current goal is SNo_eps_decr (ordsucc n) (omega_ordsucc n Ln) n (ordsuccI2 n).
Apply andI to the current goal.
An exact proof term for the current goal is L1a.
rewrite the current goal using add_SNo_com x w Lx Hw1 (from left to right).
An exact proof term for the current goal is L1a.
Let w' be given.
Assume Hw': w' {w + x|wSNoL x} {x + w|wSNoL x}.
We will prove w' < SNoCut {0} {eps_ m|mn}.
rewrite the current goal using eps_SNoCut n Ln (from right to left).
We will prove w' < eps_ n.
Apply binunionE {w + x|wSNoL x} {x + w|wSNoL x} w' Hw' to the current goal.
Assume Hw': w' {w + x|wSNoL x}.
Apply ReplE_impred (SNoL x) (λw ⇒ w + x) w' Hw' to the current goal.
Let w be given.
Assume Hw1: w SNoL x.
Assume Hw2: w' = w + x.
rewrite the current goal using Hw2 (from left to right).
An exact proof term for the current goal is andEL (w + x < eps_ n) (x + w < eps_ n) (L1 w Hw1).
Assume Hw': w' {x + w|wSNoL x}.
Apply ReplE_impred (SNoL x) (λw ⇒ x + w) w' Hw' to the current goal.
Let w be given.
Assume Hw1: w SNoL x.
Assume Hw2: w' = x + w.
rewrite the current goal using Hw2 (from left to right).
An exact proof term for the current goal is andER (w + x < eps_ n) (x + w < eps_ n) (L1 w Hw1).
We prove the intermediate claim L2: zSNoR x, eps_ n < z + x eps_ n < x + z.
Let z be given.
Assume Hz: z SNoR x.
Apply SNoR_E x Lx z Hz to the current goal.
Assume Hz1: SNo z.
rewrite the current goal using SNoLev_eps_ (ordsucc n) (omega_ordsucc n Ln) (from left to right).
Assume Hz2: SNoLev z ordsucc (ordsucc n).
Assume Hz3: x < z.
We prove the intermediate claim L2a: eps_ n < z + x.
Apply SNoLeLt_tra (eps_ n) z (z + x) (SNo_eps_ n Ln) Hz1 to the current goal.
An exact proof term for the current goal is SNo_add_SNo z x Hz1 Lx.
We will prove eps_ n z.
Apply SNo_pos_eps_Le n Hn z to the current goal.
We will prove z SNoS_ (ordsucc (ordsucc n)).
Apply SNoS_I (ordsucc (ordsucc n)) (nat_p_ordinal (ordsucc (ordsucc n)) (nat_ordsucc (ordsucc n) (nat_ordsucc n Hn))) z (SNoLev z) Hz2 to the current goal.
We will prove SNo_ (SNoLev z) z.
An exact proof term for the current goal is SNoLev_ z Hz1.
We will prove 0 < z.
Apply SNoLt_tra 0 x z SNo_0 Lx Hz1 to the current goal.
We will prove 0 < x.
An exact proof term for the current goal is SNo_eps_pos (ordsucc n) (omega_ordsucc n Ln).
We will prove x < z.
An exact proof term for the current goal is Hz3.
We will prove z < z + x.
rewrite the current goal using add_SNo_0R z Hz1 (from right to left) at position 1.
We will prove z + 0 < z + x.
Apply add_SNo_Lt2 z 0 x Hz1 SNo_0 Lx to the current goal.
We will prove 0 < x.
An exact proof term for the current goal is SNo_eps_pos (ordsucc n) (omega_ordsucc n Ln).
Apply andI to the current goal.
An exact proof term for the current goal is L2a.
rewrite the current goal using add_SNo_com x z Lx Hz1 (from left to right).
An exact proof term for the current goal is L2a.
Let z' be given.
Assume Hz': z' {z + x|zSNoR x} {x + z|zSNoR x}.
We will prove SNoCut {0} {eps_ m|mn} < z'.
rewrite the current goal using eps_SNoCut n Ln (from right to left).
We will prove eps_ n < z'.
Apply binunionE {z + x|zSNoR x} {x + z|zSNoR x} z' Hz' to the current goal.
Assume Hz': z' {z + x|zSNoR x}.
Apply ReplE_impred (SNoR x) (λz ⇒ z + x) z' Hz' to the current goal.
Let z be given.
Assume Hz1: z SNoR x.
Assume Hz2: z' = z + x.
rewrite the current goal using Hz2 (from left to right).
We will prove eps_ n < z + x.
An exact proof term for the current goal is andEL (eps_ n < z + x) (eps_ n < x + z) (L2 z Hz1).
Assume Hz': z' {x + z|zSNoR x}.
Apply ReplE_impred (SNoR x) (λz ⇒ x + z) z' Hz' to the current goal.
Let z be given.
Assume Hz1: z SNoR x.
Assume Hz2: z' = x + z.
rewrite the current goal using Hz2 (from left to right).
We will prove eps_ n < x + z.
An exact proof term for the current goal is andER (eps_ n < z + x) (eps_ n < x + z) (L2 z Hz1).
Let w be given.
Assume Hw: w {0}.
We will prove w < SNoCut ({w + x|wSNoL x} {x + w|wSNoL x}) ({z + x|zSNoR x} {x + z|zSNoR x}).
rewrite the current goal using SingE 0 w Hw (from left to right).
We will prove 0 < SNoCut ({w + x|wSNoL x} {x + w|wSNoL x}) ({z + x|zSNoR x} {x + z|zSNoR x}).
rewrite the current goal using add_SNo_eq x Lx x Lx (from right to left).
We will prove 0 < x + x.
rewrite the current goal using add_SNo_0L 0 SNo_0 (from right to left).
We will prove 0 + 0 < x + x.
An exact proof term for the current goal is add_SNo_Lt3 0 0 x x SNo_0 SNo_0 Lx Lx Lx2 Lx2.
Let z be given.
Assume Hz: z {eps_ m|mn}.
We will prove SNoCut ({w + x|wSNoL x} {x + w|wSNoL x}) ({z + x|zSNoR x} {x + z|zSNoR x}) < z.
rewrite the current goal using add_SNo_eq x Lx x Lx (from right to left).
We will prove x + x < z.
Apply ReplE_impred n eps_ z Hz to the current goal.
Let m be given.
Assume Hm1: m n.
Assume Hm2: z = eps_ m.
rewrite the current goal using Hm2 (from left to right).
We will prove x + x < eps_ m.
rewrite the current goal using IH m Hm1 (from right to left).
We will prove x + x < eps_ (ordsucc m) + eps_ (ordsucc m).
Set y to be the term eps_ (ordsucc m).
We will prove x + x < y + y.
We prove the intermediate claim Lm: m ω.
Apply nat_p_omega to the current goal.
An exact proof term for the current goal is nat_p_trans n Hn m Hm1.
We prove the intermediate claim Ly: SNo y.
An exact proof term for the current goal is SNo_eps_ (ordsucc m) (omega_ordsucc m Lm).
We prove the intermediate claim Lxy: x < y.
Apply SNo_eps_decr (ordsucc n) (omega_ordsucc n Ln) (ordsucc m) to the current goal.
We will prove ordsucc m ordsucc n.
Apply ordinal_ordsucc_In to the current goal.
We will prove ordinal n.
An exact proof term for the current goal is nat_p_ordinal n Hn.
An exact proof term for the current goal is Hm1.
An exact proof term for the current goal is add_SNo_Lt3 x x y y Lx Lx Ly Ly Lxy Lxy.
Theorem. (eps_1_half_eq1) The following is provable:
Proof:
Use transitivity with and eps_ 0.
An exact proof term for the current goal is eps_ordsucc_half_add 0 nat_0.
An exact proof term for the current goal is eps_0_1.
Theorem. (eps_1_half_eq2) The following is provable:
Proof:
rewrite the current goal using add_SNo_1_1_2 (from right to left).
We will prove (1 + 1) * eps_ 1 = 1.
rewrite the current goal using mul_SNo_distrR 1 1 (eps_ 1) SNo_1 SNo_1 SNo_eps_1 (from left to right).
We will prove 1 * eps_ 1 + 1 * eps_ 1 = 1.
rewrite the current goal using mul_SNo_oneL (eps_ 1) SNo_eps_1 (from left to right).
An exact proof term for the current goal is eps_1_half_eq1.
Theorem. (double_eps_1) The following is provable:
∀x y z, SNo xSNo ySNo zx + x = y + zx = eps_ 1 * (y + z)
Proof:
Let x, y and z be given.
Assume Hx Hy Hz H1.
We prove the intermediate claim Lyz: SNo (y + z).
An exact proof term for the current goal is SNo_add_SNo y z Hy Hz.
Apply mul_SNo_nonzero_cancel 2 x (eps_ 1 * (y + z)) SNo_2 neq_2_0 Hx (SNo_mul_SNo (eps_ 1) (y + z) (SNo_eps_ 1 (nat_p_omega 1 nat_1)) Lyz) to the current goal.
We will prove 2 * x = 2 * eps_ 1 * (y + z).
rewrite the current goal using mul_SNo_assoc 2 (eps_ 1) (y + z) SNo_2 (SNo_eps_ 1 (nat_p_omega 1 nat_1)) Lyz (from left to right).
We will prove 2 * x = (2 * eps_ 1) * (y + z).
rewrite the current goal using eps_1_half_eq2 (from left to right).
We will prove 2 * x = 1 * (y + z).
rewrite the current goal using mul_SNo_oneL (y + z) Lyz (from left to right).
We will prove 2 * x = y + z.
rewrite the current goal using add_SNo_1_1_2 (from right to left).
We will prove (1 + 1) * x = y + z.
rewrite the current goal using mul_SNo_distrR 1 1 x SNo_1 SNo_1 Hx (from left to right).
We will prove 1 * x + 1 * x = y + z.
rewrite the current goal using mul_SNo_oneL x Hx (from left to right).
We will prove x + x = y + z.
An exact proof term for the current goal is H1.
Theorem. (exp_SNo_1_bd) The following is provable:
∀x, SNo x1 x∀n, nat_p n1 x ^ n
Proof:
Let x be given.
Assume Hx Hx1.
Apply nat_ind to the current goal.
We will prove 1 x ^ 0.
rewrite the current goal using exp_SNo_nat_0 x Hx (from left to right).
Apply SNoLe_ref to the current goal.
Let n be given.
Assume Hn.
Assume IHn: 1 x ^ n.
We prove the intermediate claim Lxn: SNo (x ^ n).
An exact proof term for the current goal is SNo_exp_SNo_nat x Hx n Hn.
We will prove 1 x ^ (ordsucc n).
rewrite the current goal using exp_SNo_nat_S x Hx n Hn (from left to right).
We will prove 1 x * x ^ n.
rewrite the current goal using mul_SNo_oneL 1 SNo_1 (from right to left).
We will prove 1 * 1 x * x ^ n.
Apply nonneg_mul_SNo_Le2 1 1 x (x ^ n) SNo_1 SNo_1 Hx Lxn to the current goal.
We will prove 0 1.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is SNoLt_0_1.
We will prove 0 1.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is SNoLt_0_1.
We will prove 1 x.
An exact proof term for the current goal is Hx1.
We will prove 1 x ^ n.
An exact proof term for the current goal is IHn.
Theorem. (exp_SNo_2_bd) The following is provable:
∀n, nat_p nn < 2 ^ n
Proof:
Apply nat_ind to the current goal.
We will prove 0 < 2 ^ 0.
rewrite the current goal using exp_SNo_nat_0 2 SNo_2 (from left to right).
An exact proof term for the current goal is SNoLt_0_1.
Let n be given.
Assume Hn.
Assume IHn: n < 2 ^ n.
We prove the intermediate claim Ln: SNo n.
An exact proof term for the current goal is nat_p_SNo n Hn.
We prove the intermediate claim L2n: SNo (2 ^ n).
An exact proof term for the current goal is SNo_exp_SNo_nat 2 SNo_2 n Hn.
We will prove ordsucc n < 2 ^ (ordsucc n).
rewrite the current goal using exp_SNo_nat_S 2 SNo_2 n Hn (from left to right).
We will prove ordsucc n < 2 * 2 ^ n.
rewrite the current goal using add_SNo_1_ordsucc n (nat_p_omega n Hn) (from right to left).
We will prove n + 1 < 2 * 2 ^ n.
rewrite the current goal using add_SNo_1_1_2 (from right to left) at position 1.
We will prove n + 1 < (1 + 1) * 2 ^ n.
rewrite the current goal using mul_SNo_distrR 1 1 (2 ^ n) SNo_1 SNo_1 L2n (from left to right).
We will prove n + 1 < 1 * 2 ^ n + 1 * 2 ^ n.
rewrite the current goal using mul_SNo_oneL (2 ^ n) L2n (from left to right).
We will prove n + 1 < 2 ^ n + 2 ^ n.
Apply SNoLtLe_tra (n + 1) (2 ^ n + 1) (2 ^ n + 2 ^ n) (SNo_add_SNo n 1 Ln SNo_1) (SNo_add_SNo (2 ^ n) 1 L2n SNo_1) (SNo_add_SNo (2 ^ n) (2 ^ n) L2n L2n) to the current goal.
We will prove n + 1 < 2 ^ n + 1.
An exact proof term for the current goal is add_SNo_Lt1 n 1 (2 ^ n) Ln SNo_1 L2n IHn.
We will prove 2 ^ n + 1 2 ^ n + 2 ^ n.
Apply add_SNo_Le2 (2 ^ n) 1 (2 ^ n) L2n SNo_1 L2n to the current goal.
We will prove 1 2 ^ n.
Apply exp_SNo_1_bd 2 SNo_2 to the current goal.
We will prove 1 2.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is SNoLt_1_2.
An exact proof term for the current goal is Hn.
Theorem. (mul_SNo_eps_power_2) The following is provable:
∀n, nat_p neps_ n * 2 ^ n = 1
Proof:
Apply nat_ind to the current goal.
We will prove eps_ 0 * 2 ^ 0 = 1.
rewrite the current goal using eps_0_1 (from left to right).
rewrite the current goal using exp_SNo_nat_0 2 SNo_2 (from left to right).
We will prove 1 * 1 = 1.
An exact proof term for the current goal is mul_SNo_oneR 1 SNo_1.
Let n be given.
Assume Hn.
Assume IHn: eps_ n * 2 ^ n = 1.
We will prove eps_ (ordsucc n) * 2 ^ (ordsucc n) = 1.
rewrite the current goal using exp_SNo_nat_S 2 SNo_2 n Hn (from left to right).
We will prove eps_ (ordsucc n) * (2 * 2 ^ n) = 1.
We prove the intermediate claim LeSn: SNo (eps_ (ordsucc n)).
An exact proof term for the current goal is SNo_eps_ (ordsucc n) (nat_p_omega (ordsucc n) (nat_ordsucc n Hn)).
rewrite the current goal using mul_SNo_assoc (eps_ (ordsucc n)) 2 (2 ^ n) LeSn SNo_2 (SNo_exp_SNo_nat 2 SNo_2 n Hn) (from left to right).
We will prove (eps_ (ordsucc n) * 2) * 2 ^ n = 1.
rewrite the current goal using add_SNo_1_1_2 (from right to left) at position 1.
We will prove (eps_ (ordsucc n) * (1 + 1)) * 2 ^ n = 1.
rewrite the current goal using mul_SNo_distrL (eps_ (ordsucc n)) 1 1 LeSn SNo_1 SNo_1 (from left to right).
We will prove ((eps_ (ordsucc n) * 1) + (eps_ (ordsucc n) * 1)) * 2 ^ n = 1.
rewrite the current goal using mul_SNo_oneR (eps_ (ordsucc n)) LeSn (from left to right).
We will prove (eps_ (ordsucc n) + eps_ (ordsucc n)) * 2 ^ n = 1.
rewrite the current goal using eps_ordsucc_half_add n Hn (from left to right).
We will prove eps_ n * 2 ^ n = 1.
An exact proof term for the current goal is IHn.
Theorem. (eps_bd_1) The following is provable:
Proof:
Let n be given.
Assume Hn.
We will prove eps_ n 1.
rewrite the current goal using mul_SNo_oneR (eps_ n) (SNo_eps_ n Hn) (from right to left).
We will prove eps_ n * 1 1.
rewrite the current goal using mul_SNo_eps_power_2 n (omega_nat_p n Hn) (from right to left) at position 2.
We will prove eps_ n * 1 eps_ n * 2 ^ n.
Apply nonneg_mul_SNo_Le (eps_ n) 1 (2 ^ n) (SNo_eps_ n Hn) to the current goal.
We will prove 0 eps_ n.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is SNo_eps_pos n Hn.
An exact proof term for the current goal is SNo_1.
We will prove SNo (2 ^ n).
Apply SNo_exp_SNo_nat to the current goal.
An exact proof term for the current goal is SNo_2.
We will prove nat_p n.
An exact proof term for the current goal is omega_nat_p n Hn.
We will prove 1 2 ^ n.
Apply exp_SNo_1_bd 2 SNo_2 to the current goal.
We will prove 1 2.
Apply SNoLtLe to the current goal.
An exact proof term for the current goal is SNoLt_1_2.
We will prove nat_p n.
An exact proof term for the current goal is omega_nat_p n Hn.
Theorem. (mul_SNo_eps_power_2') The following is provable:
∀n, nat_p n2 ^ n * eps_ n = 1
Proof:
Let n be given.
Assume Hn.
Use transitivity with and eps_ n * 2 ^ n.
Apply mul_SNo_com to the current goal.
An exact proof term for the current goal is SNo_exp_SNo_nat 2 SNo_2 n Hn.
An exact proof term for the current goal is SNo_eps_ n (nat_p_omega n Hn).
An exact proof term for the current goal is mul_SNo_eps_power_2 n Hn.
Theorem. (exp_SNo_nat_mul_add) The following is provable:
∀x, SNo x∀m, nat_p m∀n, nat_p nx ^ m * x ^ n = x ^ (m + n)
Proof:
Let x be given.
Assume Hx.
Let m be given.
Assume Hm.
We prove the intermediate claim Lm: SNo m.
An exact proof term for the current goal is nat_p_SNo m Hm.
Apply nat_ind to the current goal.
We will prove x ^ m * x ^ 0 = x ^ (m + 0).
rewrite the current goal using exp_SNo_nat_0 x Hx (from left to right).
rewrite the current goal using add_SNo_0R m Lm (from left to right).
An exact proof term for the current goal is mul_SNo_oneR (x ^ m) (SNo_exp_SNo_nat x Hx m Hm).
Let n be given.
Assume Hn: nat_p n.
Assume IHn: x ^ m * x ^ n = x ^ (m + n).
We will prove x ^ m * x ^ (ordsucc n) = x ^ (m + ordsucc n).
rewrite the current goal using exp_SNo_nat_S x Hx n Hn (from left to right).
We will prove x ^ m * (x * x ^ n) = x ^ (m + ordsucc n).
rewrite the current goal using add_nat_add_SNo m (nat_p_omega m Hm) (ordsucc n) (omega_ordsucc n (nat_p_omega n Hn)) (from right to left).
We will prove x ^ m * (x * x ^ n) = x ^ (add_nat m (ordsucc n)).
rewrite the current goal using add_nat_SR m n Hn (from left to right).
We will prove x ^ m * (x * x ^ n) = x ^ (ordsucc (add_nat m n)).
rewrite the current goal using exp_SNo_nat_S x Hx (add_nat m n) (add_nat_p m Hm n Hn) (from left to right).
We will prove x ^ m * (x * x ^ n) = x * x ^ (add_nat m n).
rewrite the current goal using add_nat_add_SNo m (nat_p_omega m Hm) n (nat_p_omega n Hn) (from left to right).
We will prove x ^ m * (x * x ^ n) = x * x ^ (m + n).
rewrite the current goal using IHn (from right to left).
We will prove x ^ m * (x * x ^ n) = x * (x ^ m * x ^ n).
rewrite the current goal using mul_SNo_assoc (x ^ m) x (x ^ n) (SNo_exp_SNo_nat x Hx m Hm) Hx (SNo_exp_SNo_nat x Hx n Hn) (from left to right).
We will prove (x ^ m * x) * x ^ n = x * (x ^ m * x ^ n).
rewrite the current goal using mul_SNo_com (x ^ m) x (SNo_exp_SNo_nat x Hx m Hm) Hx (from left to right).
We will prove (x * x ^ m) * x ^ n = x * (x ^ m * x ^ n).
Use symmetry.
An exact proof term for the current goal is mul_SNo_assoc x (x ^ m) (x ^ n) Hx (SNo_exp_SNo_nat x Hx m Hm) (SNo_exp_SNo_nat x Hx n Hn).
Theorem. (exp_SNo_nat_mul_add') The following is provable:
∀x, SNo xm nω, x ^ m * x ^ n = x ^ (m + n)
Proof:
Let x be given.
Assume Hx.
Let m be given.
Assume Hm.
Let n be given.
Assume Hn.
An exact proof term for the current goal is exp_SNo_nat_mul_add x Hx m (omega_nat_p m Hm) n (omega_nat_p n Hn).
Theorem. (exp_SNo_nat_pos) The following is provable:
∀x, SNo x0 < x∀n, nat_p n0 < x ^ n
Proof:
Let x be given.
Assume Hx Hxpos.
Apply nat_ind to the current goal.
We will prove 0 < x ^ 0.
rewrite the current goal using exp_SNo_nat_0 x Hx (from left to right).
We will prove 0 < 1.
An exact proof term for the current goal is SNoLt_0_1.
Let n be given.
Assume Hn: nat_p n.
Assume IHn: 0 < x ^ n.
We will prove 0 < x ^ (ordsucc n).
rewrite the current goal using exp_SNo_nat_S x Hx n Hn (from left to right).
We will prove 0 < x * x ^ n.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left).
We will prove x * 0 < x * x ^ n.
An exact proof term for the current goal is pos_mul_SNo_Lt x 0 (x ^ n) Hx Hxpos SNo_0 (SNo_exp_SNo_nat x Hx n Hn) IHn.
Theorem. (mul_SNo_eps_eps_add_SNo) The following is provable:
m nω, eps_ m * eps_ n = eps_ (m + n)
Proof:
Let m be given.
Assume Hm.
Let n be given.
Assume Hn.
We prove the intermediate claim Lmn1: m + n ω.
An exact proof term for the current goal is add_SNo_In_omega m Hm n Hn.
We prove the intermediate claim Lmn2: nat_p (m + n).
An exact proof term for the current goal is omega_nat_p (m + n) Lmn1.
We prove the intermediate claim Lem: SNo (eps_ m).
An exact proof term for the current goal is SNo_eps_ m Hm.
We prove the intermediate claim Len: SNo (eps_ n).
An exact proof term for the current goal is SNo_eps_ n Hn.
We prove the intermediate claim Lemen: SNo (eps_ m * eps_ n).
An exact proof term for the current goal is SNo_mul_SNo (eps_ m) (eps_ n) Lem Len.
We prove the intermediate claim Lemn: SNo (eps_ (m + n)).
An exact proof term for the current goal is SNo_eps_ (m + n) Lmn1.
We prove the intermediate claim L2m: SNo (2 ^ m).
An exact proof term for the current goal is SNo_exp_SNo_nat 2 SNo_2 m (omega_nat_p m Hm).
We prove the intermediate claim L2n: SNo (2 ^ n).
An exact proof term for the current goal is SNo_exp_SNo_nat 2 SNo_2 n (omega_nat_p n Hn).
Apply mul_SNo_nonzero_cancel (2 ^ (m + n)) to the current goal.
We will prove SNo (2 ^ (m + n)).
An exact proof term for the current goal is SNo_exp_SNo_nat 2 SNo_2 (m + n) Lmn2.
We will prove 2 ^ (m + n) 0.
Assume H1: 2 ^ (m + n) = 0.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
rewrite the current goal using H1 (from right to left) at position 2.
We will prove 0 < 2 ^ (m + n).
An exact proof term for the current goal is exp_SNo_nat_pos 2 SNo_2 SNoLt_0_2 (m + n) Lmn2.
We will prove SNo (eps_ m * eps_ n).
An exact proof term for the current goal is Lemen.
We will prove SNo (eps_ (m + n)).
An exact proof term for the current goal is Lemn.
We will prove 2 ^ (m + n) * (eps_ m * eps_ n) = 2 ^ (m + n) * eps_ (m + n).
Use transitivity with (2 ^ m * 2 ^ n) * (eps_ m * eps_ n), 2 ^ m * (2 ^ n * (eps_ m * eps_ n)), 2 ^ m * eps_ m, and 1.
Use f_equal.
Use symmetry.
An exact proof term for the current goal is exp_SNo_nat_mul_add' 2 SNo_2 m Hm n Hn.
Use symmetry.
An exact proof term for the current goal is mul_SNo_assoc (2 ^ m) (2 ^ n) (eps_ m * eps_ n) L2m L2n (SNo_mul_SNo (eps_ m) (eps_ n) Lem Len).
Use f_equal.
We will prove 2 ^ n * (eps_ m * eps_ n) = eps_ m.
Use transitivity with (2 ^ n * eps_ n) * eps_ m, and 1 * eps_ m.
rewrite the current goal using mul_SNo_com (eps_ m) (eps_ n) Lem Len (from left to right).
An exact proof term for the current goal is mul_SNo_assoc (2 ^ n) (eps_ n) (eps_ m) L2n Len Lem.
Use f_equal.
An exact proof term for the current goal is mul_SNo_eps_power_2' n (omega_nat_p n Hn).
An exact proof term for the current goal is mul_SNo_oneL (eps_ m) Lem.
We will prove 2 ^ m * eps_ m = 1.
An exact proof term for the current goal is mul_SNo_eps_power_2' m (omega_nat_p m Hm).
Use symmetry.
We will prove 2 ^ (m + n) * eps_ (m + n) = 1.
An exact proof term for the current goal is mul_SNo_eps_power_2' (m + n) Lmn2.
Theorem. (SNoS_omega_Lev_equip) The following is provable:
∀n, nat_p nequip {xSNoS_ ω|SNoLev x = n} (2 ^ n)
Proof:
Apply nat_ind to the current goal.
We will prove equip {xSNoS_ ω|SNoLev x = 0} (2 ^ 0).
rewrite the current goal using exp_SNo_nat_0 2 SNo_2 (from left to right).
We will prove equip {xSNoS_ ω|SNoLev x = 0} 1.
Apply equip_sym to the current goal.
We will prove f : setset, bij 1 {xSNoS_ ω|SNoLev x = 0} f.
Set f to be the term λi : set0.
We use f to witness the existential quantifier.
Apply bijI to the current goal.
Let i be given.
Assume Hi: i 1.
We will prove 0 {xSNoS_ ω|SNoLev x = 0}.
Apply SepI to the current goal.
We will prove 0 SNoS_ ω.
Apply SNoS_I ω omega_ordinal 0 0 (nat_p_omega 0 nat_0) to the current goal.
We will prove SNo_ 0 0.
An exact proof term for the current goal is ordinal_SNo_ 0 ordinal_Empty.
We will prove SNoLev 0 = 0.
An exact proof term for the current goal is SNoLev_0.
Let i be given.
Assume Hi: i 1.
Let j be given.
Assume Hj: j 1.
Assume Hij: 0 = 0.
We will prove i = j.
Apply cases_1 i Hi to the current goal.
Apply cases_1 j Hj to the current goal.
We will prove 0 = 0.
An exact proof term for the current goal is Hij.
Let x be given.
Assume Hx: x {xSNoS_ ω|SNoLev x = 0}.
Apply SepE (SNoS_ ω) (λx ⇒ SNoLev x = 0) x Hx to the current goal.
Assume Hx1: x SNoS_ ω.
Assume Hx2: SNoLev x = 0.
Apply SNoS_E2 ω omega_ordinal x Hx1 to the current goal.
Assume Hx1a: SNoLev x ω.
Assume Hx1b: ordinal (SNoLev x).
Assume Hx1c: SNo x.
Assume Hx1d: SNo_ (SNoLev x) x.
We prove the intermediate claim L1: x = 0.
Use symmetry.
Apply SNo_eq 0 x SNo_0 Hx1c to the current goal.
We will prove SNoLev 0 = SNoLev x.
rewrite the current goal using SNoLev_0 (from left to right).
Use symmetry.
An exact proof term for the current goal is Hx2.
We will prove SNoEq_ (SNoLev 0) 0 x.
rewrite the current goal using SNoLev_0 (from left to right).
Let alpha be given.
Assume Ha: alpha 0.
We will prove False.
An exact proof term for the current goal is EmptyE alpha Ha.
We use 0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is In_0_1.
We will prove 0 = x.
Use symmetry.
An exact proof term for the current goal is L1.
Let n be given.
Assume Hn.
Assume IHn: equip {xSNoS_ ω|SNoLev x = n} (2 ^ n).
We will prove equip {xSNoS_ ω|SNoLev x = ordsucc n} (2 ^ ordsucc n).
rewrite the current goal using exp_SNo_nat_S 2 SNo_2 n Hn (from left to right).
We will prove equip {xSNoS_ ω|SNoLev x = ordsucc n} (2 * 2 ^ n).
rewrite the current goal using add_SNo_1_1_2 (from right to left) at position 1.
We will prove equip {xSNoS_ ω|SNoLev x = ordsucc n} ((1 + 1) * 2 ^ n).
rewrite the current goal using mul_SNo_distrR 1 1 (2 ^ n) SNo_1 SNo_1 (SNo_exp_SNo_nat 2 SNo_2 n Hn) (from left to right).
rewrite the current goal using mul_SNo_oneL (2 ^ n) (SNo_exp_SNo_nat 2 SNo_2 n Hn) (from left to right).
We will prove equip {xSNoS_ ω|SNoLev x = ordsucc n} (2 ^ n + 2 ^ n).
We prove the intermediate claim L2n1: nat_p (2 ^ n).
An exact proof term for the current goal is nat_exp_SNo_nat 2 nat_2 n Hn.
We prove the intermediate claim L2n2: ordinal (2 ^ n).
Apply nat_p_ordinal to the current goal.
An exact proof term for the current goal is L2n1.
We prove the intermediate claim L2n3: SNo (2 ^ n).
Apply ordinal_SNo to the current goal.
An exact proof term for the current goal is L2n2.
We prove the intermediate claim L2n2n1: nat_p (2 ^ n + 2 ^ n).
Apply omega_nat_p to the current goal.
Apply add_SNo_In_omega to the current goal.
Apply nat_p_omega to the current goal.
An exact proof term for the current goal is L2n1.
Apply nat_p_omega to the current goal.
An exact proof term for the current goal is L2n1.
We prove the intermediate claim L2n2n2: ordinal (2 ^ n + 2 ^ n).
Apply nat_p_ordinal to the current goal.
An exact proof term for the current goal is L2n2n1.
We prove the intermediate claim L2n2n3: SNo (2 ^ n + 2 ^ n).
Apply ordinal_SNo to the current goal.
An exact proof term for the current goal is L2n2n2.
We prove the intermediate claim L2npLt2n2n: ∀m, SNo mm < 2 ^ n2 ^ n + m < 2 ^ n + 2 ^ n.
Let m be given.
Assume Hm H1.
An exact proof term for the current goal is add_SNo_Lt2 (2 ^ n) m (2 ^ n) L2n3 Hm L2n3 H1.
We prove the intermediate claim L2nLt2n2n: 2 ^ n < 2 ^ n + 2 ^ n.
rewrite the current goal using add_SNo_0R (2 ^ n) L2n3 (from right to left) at position 1.
We will prove 2 ^ n + 0 < 2 ^ n + 2 ^ n.
Apply L2npLt2n2n 0 SNo_0 to the current goal.
We will prove 0 < 2 ^ n.
An exact proof term for the current goal is exp_SNo_nat_pos 2 SNo_2 SNoLt_0_2 n Hn.
Apply IHn to the current goal.
Let f be given.
Assume Hf: bij {xSNoS_ ω|SNoLev x = n} (2 ^ n) f.
Apply bijE {xSNoS_ ω|SNoLev x = n} (2 ^ n) f Hf to the current goal.
Assume Hf1: u{xSNoS_ ω|SNoLev x = n}, f u (2 ^ n).
Assume Hf2: u v{xSNoS_ ω|SNoLev x = n}, f u = f vu = v.
Assume Hf3: w2 ^ n, u{xSNoS_ ω|SNoLev x = n}, f u = w.
We prove the intermediate claim L2: x{xSNoS_ ω|SNoLev x = ordsucc n}, ∀p : prop, (SNo xSNoLev x = ordsucc n(x SNoElts_ n) {xSNoS_ ω|SNoLev x = n}SNo (x SNoElts_ n)SNoLev (x SNoElts_ n) = np)p.
Let x be given.
Assume Hx.
Apply SepE (SNoS_ ω) (λx ⇒ SNoLev x = ordsucc n) x Hx to the current goal.
Assume Hx HxSn.
Apply SNoS_E2 ω omega_ordinal x Hx to the current goal.
Assume Hx1: SNoLev x ω.
Assume Hx2: ordinal (SNoLev x).
Assume Hx3: SNo x.
Assume Hx4: SNo_ (SNoLev x) x.
Let p be given.
Assume Hp.
We prove the intermediate claim L2a: n SNoLev x.
rewrite the current goal using HxSn (from left to right).
Apply ordsuccI2 to the current goal.
We prove the intermediate claim L2b: SNoLev (x SNoElts_ n) = n.
An exact proof term for the current goal is restr_SNoLev x Hx3 n L2a.
Apply Hp to the current goal.
An exact proof term for the current goal is Hx3.
An exact proof term for the current goal is HxSn.
Apply SepI to the current goal.
We will prove (x SNoElts_ n) SNoS_ ω.
Apply SNoS_I ω omega_ordinal (x SNoElts_ n) n (nat_p_omega n Hn) to the current goal.
We will prove SNo_ n (x SNoElts_ n).
An exact proof term for the current goal is restr_SNo_ x Hx3 n L2a.
We will prove SNoLev (x SNoElts_ n) = n.
An exact proof term for the current goal is L2b.
An exact proof term for the current goal is restr_SNo x Hx3 n L2a.
An exact proof term for the current goal is L2b.
We prove the intermediate claim Lf: u{xSNoS_ ω|SNoLev x = ordsucc n}, ∀p : prop, (nat_p (f (u SNoElts_ n))ordinal (f (u SNoElts_ n))SNo (f (u SNoElts_ n))f (u SNoElts_ n) < 2 ^ np)p.
Let u be given.
Assume Hu.
Let p be given.
Assume Hp.
Apply L2 u Hu to the current goal.
Assume Hu0a: SNo u.
Assume Hu0b: SNoLev u = ordsucc n.
Assume Hu1: (u SNoElts_ n) {xSNoS_ ω|SNoLev x = n}.
Assume Hu2: SNo (u SNoElts_ n).
Assume Hu3: SNoLev (u SNoElts_ n) = n.
We prove the intermediate claim Lf1a: f (u SNoElts_ n) 2 ^ n.
An exact proof term for the current goal is Hf1 (u SNoElts_ n) Hu1.
We prove the intermediate claim Lf1b: nat_p (f (u SNoElts_ n)).
An exact proof term for the current goal is nat_p_trans (2 ^ n) L2n1 (f (u SNoElts_ n)) Lf1a.
We prove the intermediate claim Lf1c: ordinal (f (u SNoElts_ n)).
An exact proof term for the current goal is nat_p_ordinal (f (u SNoElts_ n)) Lf1b.
Apply Hp to the current goal.
An exact proof term for the current goal is Lf1b.
An exact proof term for the current goal is Lf1c.
Apply ordinal_SNo to the current goal.
An exact proof term for the current goal is Lf1c.
Apply ordinal_In_SNoLt to the current goal.
We will prove ordinal (2 ^ n).
An exact proof term for the current goal is L2n2.
We will prove f (u SNoElts_ n) 2 ^ n.
An exact proof term for the current goal is Lf1a.
We prove the intermediate claim Lg: g : setset, (∀x, n xg x = f (x SNoElts_ n)) (∀x, n xg x = 2 ^ n + f (x SNoElts_ n)).
We use (λx ⇒ if n x then f (x SNoElts_ n) else 2 ^ n + f (x SNoElts_ n)) to witness the existential quantifier.
Apply andI to the current goal.
Let x be given.
Assume H1: n x.
An exact proof term for the current goal is If_i_1 (n x) (f (x SNoElts_ n)) (2 ^ n + f (x SNoElts_ n)) H1.
Let x be given.
Assume H1: n x.
An exact proof term for the current goal is If_i_0 (n x) (f (x SNoElts_ n)) (2 ^ n + f (x SNoElts_ n)) H1.
Apply Lg to the current goal.
Let g be given.
Assume H.
Apply H to the current goal.
Assume Hg1: ∀x, n xg x = f (x SNoElts_ n).
Assume Hg2: ∀x, n xg x = 2 ^ n + f (x SNoElts_ n).
We will prove g : setset, bij {xSNoS_ ω|SNoLev x = ordsucc n} (2 ^ n + 2 ^ n) g.
We use g to witness the existential quantifier.
Apply bijI to the current goal.
Let u be given.
Assume Hu: u {xSNoS_ ω|SNoLev x = ordsucc n}.
We will prove g u 2 ^ n + 2 ^ n.
Apply L2 u Hu to the current goal.
Assume Hu0a: SNo u.
Assume Hu0b: SNoLev u = ordsucc n.
Assume Hu1: (u SNoElts_ n) {xSNoS_ ω|SNoLev x = n}.
Assume Hu2: SNo (u SNoElts_ n).
Assume Hu3: SNoLev (u SNoElts_ n) = n.
Apply Lf u Hu to the current goal.
Assume Lfu1: nat_p (f (u SNoElts_ n)).
Assume Lfu2: ordinal (f (u SNoElts_ n)).
Assume Lfu3: SNo (f (u SNoElts_ n)).
Assume Lfu4: f (u SNoElts_ n) < 2 ^ n.
Apply xm (n u) to the current goal.
Assume H1: n u.
rewrite the current goal using Hg1 u H1 (from left to right).
We will prove f (u SNoElts_ n) 2 ^ n + 2 ^ n.
Apply ordinal_SNoLt_In (f (u SNoElts_ n)) (2 ^ n + 2 ^ n) Lfu2 L2n2n2 to the current goal.
We will prove f (u SNoElts_ n) < 2 ^ n + 2 ^ n.
An exact proof term for the current goal is SNoLt_tra (f (u SNoElts_ n)) (2 ^ n) (2 ^ n + 2 ^ n) Lfu3 L2n3 L2n2n3 Lfu4 L2nLt2n2n.
Assume H1: n u.
rewrite the current goal using Hg2 u H1 (from left to right).
We will prove 2 ^ n + f (u SNoElts_ n) 2 ^ n + 2 ^ n.
Apply ordinal_SNoLt_In (2 ^ n + f (u SNoElts_ n)) (2 ^ n + 2 ^ n) (add_SNo_ordinal_ordinal (2 ^ n) L2n2 (f (u SNoElts_ n)) Lfu2) L2n2n2 to the current goal.
We will prove 2 ^ n + f (u SNoElts_ n) < 2 ^ n + 2 ^ n.
An exact proof term for the current goal is L2npLt2n2n (f (u SNoElts_ n)) Lfu3 Lfu4.
Let u be given.
Assume Hu: u {xSNoS_ ω|SNoLev x = ordsucc n}.
Let v be given.
Assume Hv: v {xSNoS_ ω|SNoLev x = ordsucc n}.
Apply L2 u Hu to the current goal.
Assume Hu0a: SNo u.
Assume Hu0b: SNoLev u = ordsucc n.
Assume Hu1: (u SNoElts_ n) {xSNoS_ ω|SNoLev x = n}.
Assume Hu2: SNo (u SNoElts_ n).
Assume Hu3: SNoLev (u SNoElts_ n) = n.
Apply Lf u Hu to the current goal.
Assume Lfu1: nat_p (f (u SNoElts_ n)).
Assume Lfu2: ordinal (f (u SNoElts_ n)).
Assume Lfu3: SNo (f (u SNoElts_ n)).
Assume Lfu4: f (u SNoElts_ n) < 2 ^ n.
Apply L2 v Hv to the current goal.
Assume Hv0a: SNo v.
Assume Hv0b: SNoLev v = ordsucc n.
Assume Hv1: (v SNoElts_ n) {xSNoS_ ω|SNoLev x = n}.
Assume Hv2: SNo (v SNoElts_ n).
Assume Hv3: SNoLev (v SNoElts_ n) = n.
Apply Lf v Hv to the current goal.
Assume Lfv1: nat_p (f (v SNoElts_ n)).
Assume Lfv2: ordinal (f (v SNoElts_ n)).
Assume Lfv3: SNo (f (v SNoElts_ n)).
Assume Lfv4: f (v SNoElts_ n) < 2 ^ n.
We prove the intermediate claim LnLu: n SNoLev u.
rewrite the current goal using Hu0b (from left to right).
Apply ordsuccI2 to the current goal.
We prove the intermediate claim LnLv: n SNoLev v.
rewrite the current goal using Hv0b (from left to right).
Apply ordsuccI2 to the current goal.
Apply xm (n u) to the current goal.
Assume H1: n u.
rewrite the current goal using Hg1 u H1 (from left to right).
Apply xm (n v) to the current goal.
Assume H2: n v.
rewrite the current goal using Hg1 v H2 (from left to right).
Assume Hguv: f (u SNoElts_ n) = f (v SNoElts_ n).
We will prove u = v.
We prove the intermediate claim Luv: u SNoElts_ n = v SNoElts_ n.
An exact proof term for the current goal is Hf2 (u SNoElts_ n) Hu1 (v SNoElts_ n) Hv1 Hguv.
Apply SNo_eq u v Hu0a Hv0a to the current goal.
We will prove SNoLev u = SNoLev v.
rewrite the current goal using Hv0b (from left to right).
An exact proof term for the current goal is Hu0b.
We will prove SNoEq_ (SNoLev u) u v.
rewrite the current goal using Hu0b (from left to right).
We will prove SNoEq_ (ordsucc n) u v.
Let i be given.
Assume Hi: i ordsucc n.
Apply ordsuccE n i Hi to the current goal.
Assume H3: i n.
We will prove i u i v.
Apply iff_trans (i u) (i u SNoElts_ n) (i v) to the current goal.
We will prove i u i u SNoElts_ n.
Apply iff_sym to the current goal.
An exact proof term for the current goal is restr_SNoEq u ?? n LnLu i H3.
We will prove i u SNoElts_ n i v.
rewrite the current goal using Luv (from left to right).
We will prove i v SNoElts_ n i v.
An exact proof term for the current goal is restr_SNoEq v ?? n LnLv i H3.
Assume H3: i = n.
rewrite the current goal using H3 (from left to right).
Apply iffI to the current goal.
Assume _.
An exact proof term for the current goal is H2.
Assume _.
An exact proof term for the current goal is H1.
Assume H2: n v.
rewrite the current goal using Hg2 v H2 (from left to right).
Assume Hguv: f (u SNoElts_ n) = 2 ^ n + f (v SNoElts_ n).
We will prove False.
Apply SNoLt_irref (2 ^ n) to the current goal.
We will prove 2 ^ n < 2 ^ n.
Apply SNoLeLt_tra (2 ^ n) (2 ^ n + f (v SNoElts_ n)) (2 ^ n) L2n3 (SNo_add_SNo (2 ^ n) (f (v SNoElts_ n)) L2n3 Lfv3) L2n3 to the current goal.
We will prove 2 ^ n 2 ^ n + f (v SNoElts_ n).
rewrite the current goal using add_SNo_0R (2 ^ n) L2n3 (from right to left) at position 1.
We will prove 2 ^ n + 0 2 ^ n + f (v SNoElts_ n).
Apply add_SNo_Le2 (2 ^ n) 0 (f (v SNoElts_ n)) L2n3 SNo_0 Lfv3 to the current goal.
We will prove 0 f (v SNoElts_ n).
An exact proof term for the current goal is omega_nonneg (f (v SNoElts_ n)) (nat_p_omega (f (v SNoElts_ n)) Lfv1).
We will prove 2 ^ n + f (v SNoElts_ n) < 2 ^ n.
rewrite the current goal using Hguv (from right to left).
An exact proof term for the current goal is Lfu4.
Assume H1: n u.
rewrite the current goal using Hg2 u H1 (from left to right).
Apply xm (n v) to the current goal.
Assume H2: n v.
rewrite the current goal using Hg1 v H2 (from left to right).
Assume Hguv: 2 ^ n + f (u SNoElts_ n) = f (v SNoElts_ n).
We will prove False.
Apply SNoLt_irref (2 ^ n) to the current goal.
We will prove 2 ^ n < 2 ^ n.
Apply SNoLeLt_tra (2 ^ n) (2 ^ n + f (u SNoElts_ n)) (2 ^ n) L2n3 (SNo_add_SNo (2 ^ n) (f (u SNoElts_ n)) L2n3 Lfu3) L2n3 to the current goal.
We will prove 2 ^ n 2 ^ n + f (u SNoElts_ n).
rewrite the current goal using add_SNo_0R (2 ^ n) L2n3 (from right to left) at position 1.
We will prove 2 ^ n + 0 2 ^ n + f (u SNoElts_ n).
Apply add_SNo_Le2 (2 ^ n) 0 (f (u SNoElts_ n)) L2n3 SNo_0 Lfu3 to the current goal.
We will prove 0 f (u SNoElts_ n).
An exact proof term for the current goal is omega_nonneg (f (u SNoElts_ n)) (nat_p_omega (f (u SNoElts_ n)) Lfu1).
We will prove 2 ^ n + f (u SNoElts_ n) < 2 ^ n.
rewrite the current goal using Hguv (from left to right).
An exact proof term for the current goal is Lfv4.
Assume H2: n v.
rewrite the current goal using Hg2 v H2 (from left to right).
Assume Hguv: 2 ^ n + f (u SNoElts_ n) = 2 ^ n + f (v SNoElts_ n).
We will prove u = v.
We prove the intermediate claim Luv: u SNoElts_ n = v SNoElts_ n.
Apply Hf2 (u SNoElts_ n) Hu1 (v SNoElts_ n) Hv1 to the current goal.
We will prove f (u SNoElts_ n) = f (v SNoElts_ n).
An exact proof term for the current goal is add_SNo_cancel_L (2 ^ n) (f (u SNoElts_ n)) (f (v SNoElts_ n)) L2n3 Lfu3 Lfv3 Hguv.
Apply SNo_eq u v Hu0a Hv0a to the current goal.
We will prove SNoLev u = SNoLev v.
rewrite the current goal using Hv0b (from left to right).
An exact proof term for the current goal is Hu0b.
We will prove SNoEq_ (SNoLev u) u v.
rewrite the current goal using Hu0b (from left to right).
We will prove SNoEq_ (ordsucc n) u v.
Let i be given.
Assume Hi: i ordsucc n.
Apply ordsuccE n i Hi to the current goal.
Assume H3: i n.
We will prove i u i v.
Apply iff_trans (i u) (i u SNoElts_ n) (i v) to the current goal.
We will prove i u i u SNoElts_ n.
Apply iff_sym to the current goal.
An exact proof term for the current goal is restr_SNoEq u ?? n LnLu i H3.
We will prove i u SNoElts_ n i v.
rewrite the current goal using Luv (from left to right).
We will prove i v SNoElts_ n i v.
An exact proof term for the current goal is restr_SNoEq v ?? n LnLv i H3.
Assume H3: i = n.
rewrite the current goal using H3 (from left to right).
Apply iffI to the current goal.
Assume H4: n u.
We will prove False.
An exact proof term for the current goal is H1 H4.
Assume H4: n v.
We will prove False.
An exact proof term for the current goal is H2 H4.
We will prove m2 ^ n + 2 ^ n, u{xSNoS_ ω|SNoLev x = ordsucc n}, g u = m.
Let m be given.
Assume Hm: m 2 ^ n + 2 ^ n.
We prove the intermediate claim Lm1: nat_p m.
An exact proof term for the current goal is nat_p_trans (2 ^ n + 2 ^ n) L2n2n1 m Hm.
We prove the intermediate claim Lm2: SNo m.
An exact proof term for the current goal is nat_p_SNo m Lm1.
Apply add_SNo_omega_In_cases m (2 ^ n) (nat_p_omega (2 ^ n) L2n1) (2 ^ n) L2n1 Hm to the current goal.
Assume H1: m 2 ^ n.
Apply Hf3 m H1 to the current goal.
Let u be given.
Assume H.
Apply H to the current goal.
Assume Hu1: u {xSNoS_ ω|SNoLev x = n}.
Assume Hu2: f u = m.
Apply SepE (SNoS_ ω) (λx ⇒ SNoLev x = n) u Hu1 to the current goal.
Assume Hu3: u SNoS_ ω.
Assume Hu4: SNoLev u = n.
Apply SNoS_E2 ω omega_ordinal u Hu3 to the current goal.
Assume Hu3a Hu3b Hu3c Hu3d.
We prove the intermediate claim Lu1: SNo (SNo_extend1 u).
An exact proof term for the current goal is SNo_extend1_SNo u ??.
We prove the intermediate claim Lu2: SNoLev (SNo_extend1 u) = ordsucc n.
rewrite the current goal using Hu4 (from right to left).
An exact proof term for the current goal is SNo_extend1_SNoLev u ??.
We prove the intermediate claim Lu3: n SNo_extend1 u.
rewrite the current goal using Hu4 (from right to left).
An exact proof term for the current goal is SNo_extend1_In u ??.
We use (SNo_extend1 u) to witness the existential quantifier.
Apply andI to the current goal.
Apply SepI to the current goal.
Apply SNoS_I ω omega_ordinal (SNo_extend1 u) (ordsucc n) to the current goal.
We will prove ordsucc n ω.
Apply omega_ordsucc to the current goal.
An exact proof term for the current goal is nat_p_omega n Hn.
We will prove SNo_ (ordsucc n) (SNo_extend1 u).
rewrite the current goal using Hu4 (from right to left).
We will prove SNo_ (ordsucc (SNoLev u)) (SNo_extend1 u).
An exact proof term for the current goal is SNo_extend1_SNo_ u ??.
We will prove SNoLev (SNo_extend1 u) = ordsucc n.
An exact proof term for the current goal is Lu2.
We will prove g (SNo_extend1 u) = m.
rewrite the current goal using Hg1 (SNo_extend1 u) Lu3 (from left to right).
We will prove f (SNo_extend1 u SNoElts_ n) = m.
rewrite the current goal using Hu4 (from right to left).
We will prove f (SNo_extend1 u SNoElts_ (SNoLev u)) = m.
rewrite the current goal using SNo_extend1_restr_eq u Hu3c (from right to left).
An exact proof term for the current goal is Hu2.
Assume H1: m + - 2 ^ n 2 ^ n.
Apply Hf3 (m + - 2 ^ n) H1 to the current goal.
Let u be given.
Assume H.
Apply H to the current goal.
Assume Hu1: u {xSNoS_ ω|SNoLev x = n}.
Assume Hu2: f u = m + - 2 ^ n.
Apply SepE (SNoS_ ω) (λx ⇒ SNoLev x = n) u Hu1 to the current goal.
Assume Hu3: u SNoS_ ω.
Assume Hu4: SNoLev u = n.
Apply SNoS_E2 ω omega_ordinal u Hu3 to the current goal.
Assume Hu3a Hu3b Hu3c Hu3d.
We prove the intermediate claim Lu1: SNo (SNo_extend0 u).
An exact proof term for the current goal is SNo_extend0_SNo u ??.
We prove the intermediate claim Lu2: SNoLev (SNo_extend0 u) = ordsucc n.
rewrite the current goal using Hu4 (from right to left).
An exact proof term for the current goal is SNo_extend0_SNoLev u ??.
We prove the intermediate claim Lu3: n SNo_extend0 u.
rewrite the current goal using Hu4 (from right to left).
An exact proof term for the current goal is SNo_extend0_nIn u ??.
We use (SNo_extend0 u) to witness the existential quantifier.
Apply andI to the current goal.
Apply SepI to the current goal.
Apply SNoS_I ω omega_ordinal (SNo_extend0 u) (ordsucc n) to the current goal.
We will prove ordsucc n ω.
Apply omega_ordsucc to the current goal.
An exact proof term for the current goal is nat_p_omega n Hn.
We will prove SNo_ (ordsucc n) (SNo_extend0 u).
rewrite the current goal using Hu4 (from right to left).
We will prove SNo_ (ordsucc (SNoLev u)) (SNo_extend0 u).
An exact proof term for the current goal is SNo_extend0_SNo_ u ??.
We will prove SNoLev (SNo_extend0 u) = ordsucc n.
An exact proof term for the current goal is Lu2.
We will prove g (SNo_extend0 u) = m.
rewrite the current goal using Hg2 (SNo_extend0 u) Lu3 (from left to right).
We will prove 2 ^ n + f (SNo_extend0 u SNoElts_ n) = m.
rewrite the current goal using Hu4 (from right to left) at position 2.
We will prove 2 ^ n + f (SNo_extend0 u SNoElts_ (SNoLev u)) = m.
rewrite the current goal using SNo_extend0_restr_eq u Hu3c (from right to left).
We will prove 2 ^ n + f u = m.
rewrite the current goal using Hu2 (from left to right).
We will prove 2 ^ n + (m + - 2 ^ n) = m.
rewrite the current goal using add_SNo_com (2 ^ n) (m + - 2 ^ n) L2n3 (SNo_add_SNo m (- 2 ^ n) Lm2 (SNo_minus_SNo (2 ^ n) L2n3)) (from left to right).
We will prove (m + - 2 ^ n) + 2 ^ n = m.
rewrite the current goal using add_SNo_assoc m (- 2 ^ n) (2 ^ n) Lm2 (SNo_minus_SNo (2 ^ n) L2n3) L2n3 (from right to left).
We will prove m + (- 2 ^ n + 2 ^ n) = m.
rewrite the current goal using add_SNo_minus_SNo_linv (2 ^ n) L2n3 (from left to right).
We will prove m + 0 = m.
An exact proof term for the current goal is add_SNo_0R m Lm2.
Theorem. (SNoS_finite) The following is provable:
Proof:
Let n be given.
Assume Hn.
We prove the intermediate claim Ln: nat_p n.
An exact proof term for the current goal is omega_nat_p n Hn.
We prove the intermediate claim Ln2: ordinal n.
An exact proof term for the current goal is nat_p_ordinal n Ln.
We prove the intermediate claim L1: SNoS_ n = in{xSNoS_ ω|SNoLev x = i}.
Apply set_ext to the current goal.
Let x be given.
Assume Hx.
Apply SNoS_E2 n Ln2 x Hx to the current goal.
Assume Hx1: SNoLev x n.
Assume Hx2: ordinal (SNoLev x).
Assume Hx3: SNo x.
Assume Hx4: SNo_ (SNoLev x) x.
Apply famunionI n (λi ⇒ {xSNoS_ ω|SNoLev x = i}) (SNoLev x) x Hx1 to the current goal.
We will prove x {x'SNoS_ ω|SNoLev x' = SNoLev x}.
Apply SepI to the current goal.
We will prove x SNoS_ ω.
Apply SNoS_Subq n ω Ln2 omega_ordinal to the current goal.
We will prove n ω.
An exact proof term for the current goal is omega_TransSet n Hn.
An exact proof term for the current goal is Hx.
We will prove SNoLev x = SNoLev x.
Use reflexivity.
Let x be given.
Assume Hx.
Apply famunionE_impred n (λi ⇒ {xSNoS_ ω|SNoLev x = i}) x Hx to the current goal.
Let i be given.
Assume Hi: i n.
Assume H1.
Apply SepE (SNoS_ ω) (λx ⇒ SNoLev x = i) x H1 to the current goal.
Assume Hxa: x SNoS_ ω.
Assume Hxb: SNoLev x = i.
We will prove x SNoS_ n.
Apply SNoS_E2 ω omega_ordinal x Hxa to the current goal.
Assume Hx1: SNoLev x ω.
Assume Hx2: ordinal (SNoLev x).
Assume Hx3: SNo x.
Assume Hx4: SNo_ (SNoLev x) x.
Apply SNoS_I n Ln2 x (SNoLev x) to the current goal.
We will prove SNoLev x n.
rewrite the current goal using Hxb (from left to right).
An exact proof term for the current goal is Hi.
We will prove SNo_ (SNoLev x) x.
An exact proof term for the current goal is Hx4.
rewrite the current goal using L1 (from left to right).
We will prove finite (in{xSNoS_ ω|SNoLev x = i}).
Apply famunion_nat_finite (λi ⇒ {xSNoS_ ω|SNoLev x = i}) n Ln to the current goal.
Let i be given.
Assume Hi: i n.
We will prove finite {xSNoS_ ω|SNoLev x = i}.
We will prove mω, equip {xSNoS_ ω|SNoLev x = i} m.
We use 2 ^ i to witness the existential quantifier.
Apply andI to the current goal.
We will prove 2 ^ i ω.
Apply nat_p_omega to the current goal.
An exact proof term for the current goal is nat_exp_SNo_nat 2 nat_2 i (nat_p_trans n Ln i Hi).
An exact proof term for the current goal is SNoS_omega_Lev_equip i (nat_p_trans n Ln i Hi).
Proof:
Let x be given.
Assume Hx.
Apply SNoS_E2 ω omega_ordinal x Hx to the current goal.
Assume Hx1: SNoLev x ω.
Assume Hx2: ordinal (SNoLev x).
Assume Hx3: SNo x.
Assume Hx4: SNo_ (SNoLev x) x.
Apply Subq_finite (SNoS_ (SNoLev x)) to the current goal.
We will prove finite (SNoS_ (SNoLev x)).
An exact proof term for the current goal is SNoS_finite (SNoLev x) Hx1.
We will prove SNoL x SNoS_ (SNoLev x).
Let y be given.
Assume Hy.
Apply SNoL_E x Hx3 y Hy to the current goal.
Assume Hy1 Hy2 Hy3.
We will prove y SNoS_ (SNoLev x).
Apply SNoS_I2 y x ?? ?? to the current goal.
We will prove SNoLev y SNoLev x.
An exact proof term for the current goal is Hy2.
Proof:
Let x be given.
Assume Hx.
Apply SNoS_E2 ω omega_ordinal x Hx to the current goal.
Assume Hx1: SNoLev x ω.
Assume Hx2: ordinal (SNoLev x).
Assume Hx3: SNo x.
Assume Hx4: SNo_ (SNoLev x) x.
Apply Subq_finite (SNoS_ (SNoLev x)) to the current goal.
We will prove finite (SNoS_ (SNoLev x)).
An exact proof term for the current goal is SNoS_finite (SNoLev x) Hx1.
We will prove SNoR x SNoS_ (SNoLev x).
Let y be given.
Assume Hy.
Apply SNoR_E x Hx3 y Hy to the current goal.
Assume Hy1 Hy2 Hy3.
We will prove y SNoS_ (SNoLev x).
Apply SNoS_I2 y x ?? ?? to the current goal.
We will prove SNoLev y SNoLev x.
An exact proof term for the current goal is Hy2.
End of Section SurrealExp