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
Axiom. (famunion_ext) We take the following as an axiom:
∀X, ∀f g : setset, (xX, f x = g x)famunion X f = famunion X g
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. (SNoL_SNoCutP_ex) We take the following as an axiom:
∀L R, SNoCutP L RwSNoL (SNoCut L R), w'L, w w'
Axiom. (SNoR_SNoCutP_ex) We take the following as an axiom:
∀L R, SNoCutP L RzSNoR (SNoCut L R), z'R, z' z
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
Axiom. (pos_low_eq_one) We take the following as an axiom:
∀x, SNo x0 < xSNoLev x 1x = 1
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_prop2) We take the following as an axiom:
∀x y, SNo xSNo yx + - x + y = y
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
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.
Axiom. (mul_SNo_eq) We take the following as an axiom:
∀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})
Axiom. (mul_SNo_eq_2) We take the following as an axiom:
∀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
Axiom. (mul_SNo_prop_1) We take the following as an axiom:
∀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
Axiom. (SNo_mul_SNo) We take the following as an axiom:
∀x y, SNo xSNo ySNo (x * y)
Axiom. (SNo_mul_SNo_lem) We take the following as an axiom:
∀x y u v, SNo xSNo ySNo uSNo vSNo (u * y + x * v + - (u * v))
Axiom. (SNo_mul_SNo_3) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zSNo (x * y * z)
Axiom. (mul_SNo_eq_3) We take the following as an axiom:
∀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
Axiom. (mul_SNo_Lt) We take the following as an axiom:
∀x y u v, SNo xSNo ySNo uSNo vu < xv < yu * y + x * v < x * y + u * v
Axiom. (mul_SNo_Le) We take the following as an axiom:
∀x y u v, SNo xSNo ySNo uSNo vu xv yu * y + x * v x * y + u * v
Axiom. (mul_SNo_SNoL_interpolate) We take the following as an axiom:
∀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)
Axiom. (mul_SNo_SNoL_interpolate_impred) We take the following as an axiom:
∀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
Axiom. (mul_SNo_SNoR_interpolate) We take the following as an axiom:
∀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)
Axiom. (mul_SNo_SNoR_interpolate_impred) We take the following as an axiom:
∀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
Axiom. (mul_SNo_Subq_lem) We take the following as an axiom:
∀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'
Axiom. (mul_SNo_zeroR) We take the following as an axiom:
∀x, SNo xx * 0 = 0
Axiom. (mul_SNo_oneR) We take the following as an axiom:
∀x, SNo xx * 1 = x
Axiom. (mul_SNo_com) We take the following as an axiom:
∀x y, SNo xSNo yx * y = y * x
Axiom. (mul_SNo_minus_distrL) We take the following as an axiom:
∀x y, SNo xSNo y(- x) * y = - x * y
Axiom. (mul_SNo_minus_distrR) We take the following as an axiom:
∀x y, SNo xSNo yx * (- y) = - (x * y)
Axiom. (mul_SNo_distrR) We take the following as an axiom:
∀x y z, SNo xSNo ySNo z(x + y) * z = x * z + y * z
Axiom. (mul_SNo_distrL) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx * (y + z) = x * y + x * z
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
Axiom. (mul_SNo_assoc_lem1) We take the following as an axiom:
∀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
Axiom. (mul_SNo_assoc_lem2) We take the following as an axiom:
∀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
End of Section mul_SNo_assoc_lems
Axiom. (mul_SNo_assoc) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx * (y * z) = (x * y) * z
Axiom. (mul_nat_mul_SNo) We take the following as an axiom:
Axiom. (mul_SNo_In_omega) We take the following as an axiom:
Axiom. (mul_SNo_zeroL) We take the following as an axiom:
∀x, SNo x0 * x = 0
Axiom. (mul_SNo_oneL) We take the following as an axiom:
∀x, SNo x1 * x = x
Axiom. (pos_mul_SNo_Lt) We take the following as an axiom:
∀x y z, SNo x0 < xSNo ySNo zy < zx * y < x * z
Axiom. (nonneg_mul_SNo_Le) We take the following as an axiom:
∀x y z, SNo x0 xSNo ySNo zy zx * y x * z
Axiom. (neg_mul_SNo_Lt) We take the following as an axiom:
∀x y z, SNo xx < 0SNo ySNo zz < yx * y < x * z
Axiom. (pos_mul_SNo_Lt') We take the following as an axiom:
∀x y z, SNo xSNo ySNo z0 < zx < yx * z < y * z
Axiom. (mul_SNo_Lt1_pos_Lt) We take the following as an axiom:
∀x y, SNo xSNo yx < 10 < yx * y < y
Axiom. (nonneg_mul_SNo_Le') We take the following as an axiom:
∀x y z, SNo xSNo ySNo z0 zx yx * z y * z
Axiom. (mul_SNo_Le1_nonneg_Le) We take the following as an axiom:
∀x y, SNo xSNo yx 10 yx * y y
Axiom. (pos_mul_SNo_Lt2) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w0 < x0 < yx < zy < wx * y < z * w
Axiom. (nonneg_mul_SNo_Le2) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w0 x0 yx zy wx * y z * w
Axiom. (mul_SNo_pos_pos) We take the following as an axiom:
∀x y, SNo xSNo y0 < x0 < y0 < x * y
Axiom. (mul_SNo_pos_neg) We take the following as an axiom:
∀x y, SNo xSNo y0 < xy < 0x * y < 0
Axiom. (mul_SNo_neg_pos) We take the following as an axiom:
∀x y, SNo xSNo yx < 00 < yx * y < 0
Axiom. (mul_SNo_neg_neg) We take the following as an axiom:
∀x y, SNo xSNo yx < 0y < 00 < x * y
Axiom. (mul_SNo_nonneg_nonneg) We take the following as an axiom:
∀x y, SNo xSNo y0 x0 y0 x * y
Axiom. (mul_SNo_nonpos_pos) We take the following as an axiom:
∀x y, SNo xSNo yx 00 < yx * y 0
Axiom. (mul_SNo_nonpos_neg) We take the following as an axiom:
∀x y, SNo xSNo yx 0y < 00 x * y
Axiom. (nonpos_mul_SNo_Le) We take the following as an axiom:
∀x y z, SNo xx 0SNo ySNo zz yx * y x * z
Axiom. (SNo_sqr_nonneg) We take the following as an axiom:
∀x, SNo x0 x * x
Axiom. (SNo_zero_or_sqr_pos) We take the following as an axiom:
∀x, SNo xx = 0 0 < x * x
Axiom. (SNo_foil) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w(x + y) * (z + w) = x * z + x * w + y * z + y * w
Axiom. (mul_SNo_minus_minus) We take the following as an axiom:
∀x y, SNo xSNo y(- x) * (- y) = x * y
Axiom. (mul_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. (mul_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. (mul_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. (mul_SNo_rotate_3_1) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx * y * z = z * x * y
Axiom. (mul_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. (SNo_foil_mm) We take the following as an axiom:
∀x y z w, SNo xSNo ySNo zSNo w(x + - y) * (z + - w) = x * z + - x * w + - y * z + y * w
Axiom. (mul_SNo_nonzero_cancel) We take the following as an axiom:
∀x y z, SNo xx 0SNo ySNo zx * y = x * zy = z
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.
Axiom. (exp_SNo_nat_0) We take the following as an axiom:
∀x, SNo xx ^ 0 = 1
Axiom. (exp_SNo_nat_S) We take the following as an axiom:
∀x, SNo x∀n, nat_p nx ^ (ordsucc n) = x * x ^ n
Axiom. (exp_SNo_nat_1) We take the following as an axiom:
∀x, SNo xx ^ 1 = x
Axiom. (SNo_exp_SNo_nat) We take the following as an axiom:
∀x, SNo x∀n, nat_p nSNo (x ^ n)
Axiom. (nat_exp_SNo_nat) We take the following as an axiom:
∀x, nat_p x∀n, nat_p nnat_p (x ^ n)
Axiom. (eps_ordsucc_half_add) We take the following as an axiom:
∀n, nat_p neps_ (ordsucc n) + eps_ (ordsucc n) = eps_ n
Axiom. (eps_1_half_eq1) We take the following as an axiom:
Axiom. (eps_1_half_eq2) We take the following as an axiom:
Axiom. (double_eps_1) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zx + x = y + zx = eps_ 1 * (y + z)
Axiom. (exp_SNo_1_bd) We take the following as an axiom:
∀x, SNo x1 x∀n, nat_p n1 x ^ n
Axiom. (exp_SNo_2_bd) We take the following as an axiom:
∀n, nat_p nn < 2 ^ n
Axiom. (mul_SNo_eps_power_2) We take the following as an axiom:
∀n, nat_p neps_ n * 2 ^ n = 1
Axiom. (eps_bd_1) We take the following as an axiom:
Axiom. (mul_SNo_eps_power_2') We take the following as an axiom:
∀n, nat_p n2 ^ n * eps_ n = 1
Axiom. (exp_SNo_nat_mul_add) We take the following as an axiom:
∀x, SNo x∀m, nat_p m∀n, nat_p nx ^ m * x ^ n = x ^ (m + n)
Axiom. (exp_SNo_nat_mul_add') We take the following as an axiom:
∀x, SNo xm nω, x ^ m * x ^ n = x ^ (m + n)
Axiom. (exp_SNo_nat_pos) We take the following as an axiom:
∀x, SNo x0 < x∀n, nat_p n0 < x ^ n
Axiom. (mul_SNo_eps_eps_add_SNo) We take the following as an axiom:
m nω, eps_ m * eps_ n = eps_ (m + n)
Axiom. (SNoS_omega_Lev_equip) We take the following as an axiom:
∀n, nat_p nequip {xSNoS_ ω|SNoLev x = n} (2 ^ n)
Axiom. (SNoS_finite) We take the following as an axiom:
Axiom. (SNoS_omega_SNoL_finite) We take the following as an axiom:
Axiom. (SNoS_omega_SNoR_finite) We take the following as an axiom:
End of Section SurrealExp
Beginning of Section Int
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.
Primitive. The name int is a term of type set.
Axiom. (int_SNo_cases) We take the following as an axiom:
∀p : setprop, (nω, p n)(nω, p (- n))xint, p x
Axiom. (int_SNo) We take the following as an axiom:
Axiom. (Subq_omega_int) We take the following as an axiom:
Axiom. (int_minus_SNo_omega) We take the following as an axiom:
Axiom. (int_add_SNo_lem) We take the following as an axiom:
nω, ∀m, nat_p m- n + m int
Axiom. (int_add_SNo) We take the following as an axiom:
Axiom. (int_minus_SNo) We take the following as an axiom:
Axiom. (int_mul_SNo) We take the following as an axiom:
End of Section Int
Beginning of Section SurrealAbs
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 abs_SNo to be λx ⇒ if 0 x then x else - x of type setset.
Axiom. (nonneg_abs_SNo) We take the following as an axiom:
∀x, 0 xabs_SNo x = x
Axiom. (not_nonneg_abs_SNo) We take the following as an axiom:
∀x, ¬ (0 x)abs_SNo x = - x
Axiom. (abs_SNo_0) We take the following as an axiom:
Axiom. (pos_abs_SNo) We take the following as an axiom:
∀x, 0 < xabs_SNo x = x
Axiom. (neg_abs_SNo) We take the following as an axiom:
∀x, SNo xx < 0abs_SNo x = - x
Axiom. (SNo_abs_SNo) We take the following as an axiom:
∀x, SNo xSNo (abs_SNo x)
Axiom. (abs_SNo_Lev) We take the following as an axiom:
∀x, SNo xSNoLev (abs_SNo x) = SNoLev x
Axiom. (abs_SNo_minus) We take the following as an axiom:
∀x, SNo xabs_SNo (- x) = abs_SNo x
Axiom. (abs_SNo_dist_swap) We take the following as an axiom:
∀x y, SNo xSNo yabs_SNo (x + - y) = abs_SNo (y + - x)
Axiom. (SNo_triangle) We take the following as an axiom:
∀x y, SNo xSNo yabs_SNo (x + y) abs_SNo x + abs_SNo y
Axiom. (SNo_triangle2) We take the following as an axiom:
∀x y z, SNo xSNo ySNo zabs_SNo (x + - z) abs_SNo (x + - y) + abs_SNo (y + - z)
End of Section SurrealAbs
Beginning of Section SNoMaxMin
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.
Notation. We use ^ as an infix operator with priority 342 and which associates to the right corresponding to applying term exp_SNo_nat.
Notation. We use < as an infix operator with priority 490 and no associativity corresponding to applying term SNoLt.
Notation. We use as an infix operator with priority 490 and no associativity corresponding to applying term SNoLe.
Definition. We define SNo_max_of to be λX x ⇒ x X SNo x yX, SNo yy x of type setsetprop.
Definition. We define SNo_min_of to be λX x ⇒ x X SNo x yX, SNo yx y of type setsetprop.
Axiom. (minus_SNo_max_min) We take the following as an axiom:
∀X y, (xX, SNo x)SNo_max_of X ySNo_min_of {- x|xX} (- y)
Axiom. (minus_SNo_max_min') We take the following as an axiom:
∀X y, (xX, SNo x)SNo_max_of {- x|xX} ySNo_min_of X (- y)
Axiom. (minus_SNo_min_max) We take the following as an axiom:
∀X y, (xX, SNo x)SNo_min_of X ySNo_max_of {- x|xX} (- y)
Axiom. (double_SNo_max_1) We take the following as an axiom:
∀x y, SNo xSNo_max_of (SNoL x) y∀z, SNo zx < zy + z < x + xwSNoR z, y + w = x + x
Axiom. (double_SNo_min_1) We take the following as an axiom:
∀x y, SNo xSNo_min_of (SNoR x) y∀z, SNo zz < xx + x < y + zwSNoL z, y + w = x + x
Axiom. (finite_max_exists) We take the following as an axiom:
∀X, (xX, SNo x)finite XX 0x, SNo_max_of X x
Axiom. (finite_min_exists) We take the following as an axiom:
∀X, (xX, SNo x)finite XX 0x, SNo_min_of X x
Axiom. (SNoS_omega_SNoL_max_exists) We take the following as an axiom:
Axiom. (SNoS_omega_SNoR_min_exists) We take the following as an axiom:
End of Section SNoMaxMin
Beginning of Section DiadicRationals
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.
Notation. We use < as an infix operator with priority 490 and no associativity corresponding to applying term SNoLt.
Notation. We use as an infix operator with priority 490 and no associativity corresponding to applying term SNoLe.
Notation. We use ^ as an infix operator with priority 342 and which associates to the right corresponding to applying term exp_SNo_nat.
Axiom. (nonneg_diadic_rational_p_SNoS_omega) We take the following as an axiom:
kω, ∀n, nat_p neps_ k * n SNoS_ ω
Definition. We define diadic_rational_p to be λx ⇒ kω, mint, x = eps_ k * m of type setprop.
Axiom. (diadic_rational_p_SNoS_omega) We take the following as an axiom:
Axiom. (int_diadic_rational_p) We take the following as an axiom:
Axiom. (omega_diadic_rational_p) We take the following as an axiom:
Axiom. (eps_diadic_rational_p) We take the following as an axiom:
Axiom. (minus_SNo_diadic_rational_p) We take the following as an axiom:
Axiom. (mul_SNo_diadic_rational_p) We take the following as an axiom:
Axiom. (add_SNo_diadic_rational_p) We take the following as an axiom:
Axiom. (SNoS_omega_diadic_rational_p_lem) We take the following as an axiom:
∀n, nat_p n∀x, SNo xSNoLev x = ndiadic_rational_p x
Axiom. (SNoS_omega_diadic_rational_p) We take the following as an axiom:
Axiom. (mul_SNo_SNoS_omega) We take the following as an axiom:
End of Section DiadicRationals
Beginning of Section SurrealDiv
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.
Notation. We use < as an infix operator with priority 490 and no associativity corresponding to applying term SNoLt.
Notation. We use as an infix operator with priority 490 and no associativity corresponding to applying term SNoLe.
Notation. We use ^ as an infix operator with priority 342 and which associates to the right corresponding to applying term exp_SNo_nat.
Definition. We define SNoL_pos to be λx ⇒ {wSNoL x|0 < w} of type setset.
Theorem. (SNo_recip_pos_pos) The following is provable:
∀x xi, SNo xSNo xi0 < xx * xi = 10 < xi
Proof:
Let x and xi be given.
Assume Hx Hxi Hxpos Hxxi.
Apply SNoLt_trichotomy_or_impred 0 xi SNo_0 Hxi to the current goal.
Assume H1.
An exact proof term for the current goal is H1.
Assume H1: 0 = xi.
We will prove False.
Apply neq_0_1 to the current goal.
We will prove 0 = 1.
rewrite the current goal using Hxxi (from right to left).
We will prove 0 = x * xi.
rewrite the current goal using H1 (from right to left).
We will prove 0 = x * 0.
Use symmetry.
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
Assume H1: xi < 0.
We will prove False.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
Apply SNoLt_tra 0 1 0 SNo_0 SNo_1 SNo_0 SNoLt_0_1 to the current goal.
We will prove 1 < 0.
rewrite the current goal using Hxxi (from right to left).
We will prove x * xi < 0.
An exact proof term for the current goal is mul_SNo_pos_neg x xi Hx Hxi Hxpos H1.
Theorem. (SNo_recip_lem1) The following is provable:
∀x x' x'i y y', SNo x0 < xx' SNoL_pos xSNo x'ix' * x'i = 1SNo yx * y < 1SNo y'1 + - x * y' = (1 + - x * y) * (x' + - x) * x'i1 < x * y'
Proof:
Let x, x', x'i, y and y' be given.
Assume Hx Hxpos Hx' Hx'i Hx'x'i Hy Hxy1 Hy' Hy'y.
Apply SepE (SNoL x) (λw ⇒ 0 < w) x' Hx' to the current goal.
Assume Hx'L: x' SNoL x.
Assume Hx'pos: 0 < x'.
Apply SNoL_E x Hx x' Hx'L to the current goal.
Assume Hx'1: SNo x'.
Assume Hx'2: SNoLev x' SNoLev x.
Assume Hx'3: x' < x.
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 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 L1: 0 < 1 + - x * y.
Apply add_SNo_minus_Lt2b 1 (x * y) 0 SNo_1 Lxy SNo_0 to the current goal.
We will prove 0 + x * y < 1.
rewrite the current goal using add_SNo_0L (x * y) Lxy (from left to right).
We will prove x * y < 1.
An exact proof term for the current goal is Hxy1.
We prove the intermediate claim L2: (x' + - x) * x'i < 0.
Apply mul_SNo_neg_pos (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i to the current goal.
We will prove x' + - x < 0.
Apply add_SNo_minus_Lt1b x' x 0 Hx'1 Hx SNo_0 to the current goal.
We will prove x' < 0 + x.
rewrite the current goal using add_SNo_0L x Hx (from left to right).
An exact proof term for the current goal is Hx'3.
We will prove 0 < x'i.
An exact proof term for the current goal is SNo_recip_pos_pos x' x'i Hx'1 Hx'i Hx'pos Hx'x'i.
We prove the intermediate claim L3: 1 + - x * y' < 0.
rewrite the current goal using Hy'y (from left to right).
We will prove (1 + - x * y) * (x' + - x) * x'i < 0.
Apply mul_SNo_pos_neg (1 + - x * y) ((x' + - x) * x'i) (SNo_add_SNo 1 (- x * y) SNo_1 (SNo_minus_SNo (x * y) Lxy)) (SNo_mul_SNo (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i) to the current goal.
An exact proof term for the current goal is L1.
An exact proof term for the current goal is L2.
We will prove 1 < x * y'.
rewrite the current goal using add_SNo_0L (x * y') Lxy' (from right to left).
We will prove 1 < 0 + x * y'.
An exact proof term for the current goal is add_SNo_minus_Lt1 1 (x * y') 0 SNo_1 Lxy' SNo_0 L3.
Theorem. (SNo_recip_lem2) The following is provable:
∀x x' x'i y y', SNo x0 < xx' SNoL_pos xSNo x'ix' * x'i = 1SNo y1 < x * ySNo y'1 + - x * y' = (1 + - x * y) * (x' + - x) * x'ix * y' < 1
Proof:
Let x, x', x'i, y and y' be given.
Assume Hx Hxpos Hx' Hx'i Hx'x'i Hy Hxy1 Hy' Hy'y.
Apply SepE (SNoL x) (λw ⇒ 0 < w) x' Hx' to the current goal.
Assume Hx'L: x' SNoL x.
Assume Hx'pos: 0 < x'.
Apply SNoL_E x Hx x' Hx'L to the current goal.
Assume Hx'1: SNo x'.
Assume Hx'2: SNoLev x' SNoLev x.
Assume Hx'3: x' < x.
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 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 L1: 1 + - x * y < 0.
Apply add_SNo_minus_Lt1b 1 (x * y) 0 SNo_1 Lxy SNo_0 to the current goal.
We will prove 1 < 0 + x * y.
rewrite the current goal using add_SNo_0L (x * y) Lxy (from left to right).
We will prove 1 < x * y.
An exact proof term for the current goal is Hxy1.
We prove the intermediate claim L2: (x' + - x) * x'i < 0.
Apply mul_SNo_neg_pos (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i to the current goal.
We will prove x' + - x < 0.
Apply add_SNo_minus_Lt1b x' x 0 Hx'1 Hx SNo_0 to the current goal.
We will prove x' < 0 + x.
rewrite the current goal using add_SNo_0L x Hx (from left to right).
An exact proof term for the current goal is Hx'3.
We will prove 0 < x'i.
An exact proof term for the current goal is SNo_recip_pos_pos x' x'i Hx'1 Hx'i Hx'pos Hx'x'i.
We prove the intermediate claim L3: 0 < 1 + - x * y'.
rewrite the current goal using Hy'y (from left to right).
We will prove 0 < (1 + - x * y) * (x' + - x) * x'i.
Apply mul_SNo_neg_neg (1 + - x * y) ((x' + - x) * x'i) (SNo_add_SNo 1 (- x * y) SNo_1 (SNo_minus_SNo (x * y) Lxy)) (SNo_mul_SNo (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i) to the current goal.
An exact proof term for the current goal is L1.
An exact proof term for the current goal is L2.
We will prove x * y' < 1.
rewrite the current goal using add_SNo_0L (x * y') Lxy' (from right to left).
We will prove 0 + x * y' < 1.
An exact proof term for the current goal is add_SNo_minus_Lt2 1 (x * y') 0 SNo_1 Lxy' SNo_0 L3.
Theorem. (SNo_recip_lem3) The following is provable:
∀x x' x'i y y', SNo x0 < xx' SNoR xSNo x'ix' * x'i = 1SNo yx * y < 1SNo y'1 + - x * y' = (1 + - x * y) * (x' + - x) * x'ix * y' < 1
Proof:
Let x, x', x'i, y and y' be given.
Assume Hx Hxpos Hx' Hx'i Hx'x'i Hy Hxy1 Hy' Hy'y.
Apply SNoR_E x Hx x' Hx' to the current goal.
Assume Hx'1: SNo x'.
Assume Hx'2: SNoLev x' SNoLev x.
Assume Hx'3: x < x'.
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 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 Lx'pos: 0 < x'.
An exact proof term for the current goal is SNoLt_tra 0 x x' SNo_0 Hx Hx'1 Hxpos Hx'3.
We prove the intermediate claim L1: 0 < 1 + - x * y.
Apply add_SNo_minus_Lt2b 1 (x * y) 0 SNo_1 Lxy SNo_0 to the current goal.
We will prove 0 + x * y < 1.
rewrite the current goal using add_SNo_0L (x * y) Lxy (from left to right).
We will prove x * y < 1.
An exact proof term for the current goal is Hxy1.
We prove the intermediate claim L2: 0 < (x' + - x) * x'i.
Apply mul_SNo_pos_pos (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i to the current goal.
We will prove 0 < x' + - x.
Apply add_SNo_minus_Lt2b x' x 0 Hx'1 Hx SNo_0 to the current goal.
We will prove 0 + x < x'.
rewrite the current goal using add_SNo_0L x Hx (from left to right).
An exact proof term for the current goal is Hx'3.
We will prove 0 < x'i.
An exact proof term for the current goal is SNo_recip_pos_pos x' x'i Hx'1 Hx'i Lx'pos Hx'x'i.
We prove the intermediate claim L3: 0 < 1 + - x * y'.
rewrite the current goal using Hy'y (from left to right).
We will prove 0 < (1 + - x * y) * (x' + - x) * x'i.
Apply mul_SNo_pos_pos (1 + - x * y) ((x' + - x) * x'i) (SNo_add_SNo 1 (- x * y) SNo_1 (SNo_minus_SNo (x * y) Lxy)) (SNo_mul_SNo (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i) to the current goal.
An exact proof term for the current goal is L1.
An exact proof term for the current goal is L2.
We will prove x * y' < 1.
rewrite the current goal using add_SNo_0L (x * y') Lxy' (from right to left).
We will prove 0 + x * y' < 1.
An exact proof term for the current goal is add_SNo_minus_Lt2 1 (x * y') 0 SNo_1 Lxy' SNo_0 L3.
Theorem. (SNo_recip_lem4) The following is provable:
∀x x' x'i y y', SNo x0 < xx' SNoR xSNo x'ix' * x'i = 1SNo y1 < x * ySNo y'1 + - x * y' = (1 + - x * y) * (x' + - x) * x'i1 < x * y'
Proof:
Let x, x', x'i, y and y' be given.
Assume Hx Hxpos Hx' Hx'i Hx'x'i Hy Hxy1 Hy' Hy'y.
Apply SNoR_E x Hx x' Hx' to the current goal.
Assume Hx'1: SNo x'.
Assume Hx'2: SNoLev x' SNoLev x.
Assume Hx'3: x < x'.
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 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 Lx'pos: 0 < x'.
An exact proof term for the current goal is SNoLt_tra 0 x x' SNo_0 Hx Hx'1 Hxpos Hx'3.
We prove the intermediate claim L1: 1 + - x * y < 0.
Apply add_SNo_minus_Lt1b 1 (x * y) 0 SNo_1 Lxy SNo_0 to the current goal.
We will prove 1 < 0 + x * y.
rewrite the current goal using add_SNo_0L (x * y) Lxy (from left to right).
We will prove 1 < x * y.
An exact proof term for the current goal is Hxy1.
We prove the intermediate claim L2: 0 < (x' + - x) * x'i.
Apply mul_SNo_pos_pos (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i to the current goal.
We will prove 0 < x' + - x.
Apply add_SNo_minus_Lt2b x' x 0 Hx'1 Hx SNo_0 to the current goal.
We will prove 0 + x < x'.
rewrite the current goal using add_SNo_0L x Hx (from left to right).
An exact proof term for the current goal is Hx'3.
We will prove 0 < x'i.
An exact proof term for the current goal is SNo_recip_pos_pos x' x'i Hx'1 Hx'i Lx'pos Hx'x'i.
We prove the intermediate claim L3: 1 + - x * y' < 0.
rewrite the current goal using Hy'y (from left to right).
We will prove (1 + - x * y) * (x' + - x) * x'i < 0.
Apply mul_SNo_neg_pos (1 + - x * y) ((x' + - x) * x'i) (SNo_add_SNo 1 (- x * y) SNo_1 (SNo_minus_SNo (x * y) Lxy)) (SNo_mul_SNo (x' + - x) x'i (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hx'i) to the current goal.
An exact proof term for the current goal is L1.
An exact proof term for the current goal is L2.
We will prove 1 < x * y'.
rewrite the current goal using add_SNo_0L (x * y') Lxy' (from right to left).
We will prove 1 < 0 + x * y'.
An exact proof term for the current goal is add_SNo_minus_Lt1 1 (x * y') 0 SNo_1 Lxy' SNo_0 L3.
Definition. We define SNo_recipauxset to be λY x X g ⇒ yY{(1 + (x' + - x) * y) * g x'|x'X} of type setsetset(setset)set.
Theorem. (SNo_recipauxset_I) The following is provable:
∀Y x X, ∀g : setset, yY, x'X, (1 + (x' + - x) * y) * g x' SNo_recipauxset Y x X g
Proof:
Let Y, x, X, g and y be given.
Assume Hy.
Let x' be given.
Assume Hx'.
We will prove (1 + (x' + - x) * y) * g x' yY{(1 + (x' + - x) * y) * g x'|x'X}.
Apply famunionI Y (λy ⇒ {(1 + (x' + - x) * y) * g x'|x'X}) y ((1 + (x' + - x) * y) * g x') Hy to the current goal.
We will prove (1 + (x' + - x) * y) * g x' {(1 + (x' + - x) * y) * g x'|x'X}.
An exact proof term for the current goal is ReplI X (λx' ⇒ (1 + (x' + - x) * y) * g x') x' Hx'.
Theorem. (SNo_recipauxset_E) The following is provable:
∀Y x X, ∀g : setset, zSNo_recipauxset Y x X g, ∀p : prop, (yY, x'X, z = (1 + (x' + - x) * y) * g x'p)p
Proof:
Let Y, x, X, g and z be given.
Assume Hz.
Let p be given.
Assume H1.
Apply famunionE_impred Y (λy ⇒ {(1 + (x' + - x) * y) * g x'|x'X}) z Hz to the current goal.
Let y be given.
Assume Hy.
Assume H2: z {(1 + (x' + - x) * y) * g x'|x'X}.
Apply ReplE_impred X (λx' ⇒ (1 + (x' + - x) * y) * g x') z H2 to the current goal.
Let x' be given.
Assume Hx': x' X.
Assume H3: z = (1 + (x' + - x) * y) * g x'.
An exact proof term for the current goal is H1 y Hy x' Hx' H3.
Theorem. (SNo_recipauxset_ext) The following is provable:
∀Y x X, ∀g h : setset, (x'X, g x' = h x')SNo_recipauxset Y x X g = SNo_recipauxset Y x X h
Proof:
Let Y, x, X, g and h be given.
Assume Hgh.
We will prove (yY{(1 + (x' + - x) * y) * g x'|x'X}) = (yY{(1 + (x' + - x) * y) * h x'|x'X}).
Apply famunion_ext to the current goal.
Let y be given.
Assume Hy.
Apply ReplEq_ext X (λx' ⇒ (1 + (x' + - x) * y) * g x') (λx' ⇒ (1 + (x' + - x) * y) * h x') to the current goal.
Let x' be given.
Assume Hx'.
We will prove (1 + (x' + - x) * y) * g x' = (1 + (x' + - x) * y) * h x'.
Use f_equal.
An exact proof term for the current goal is Hgh x' Hx'.
Definition. We define SNo_recipaux to be λx g ⇒ nat_primrec ({0},0) (λk p ⇒ (p 0 SNo_recipauxset (p 0) x (SNoR x) g SNo_recipauxset (p 1) x (SNoL_pos x) g,p 1 SNo_recipauxset (p 0) x (SNoL_pos x) g SNo_recipauxset (p 1) x (SNoR x) g)) of type set(setset)setset.
Theorem. (SNo_recipaux_0) The following is provable:
∀x, ∀g : setset, SNo_recipaux x g 0 = ({0},0)
Proof:
Let x and g be given.
An exact proof term for the current goal is nat_primrec_0 ({0},0) (λk p ⇒ (p 0 SNo_recipauxset (p 0) x (SNoR x) g SNo_recipauxset (p 1) x (SNoL_pos x) g,p 1 SNo_recipauxset (p 0) x (SNoL_pos x) g SNo_recipauxset (p 1) x (SNoR x) g)).
Theorem. (SNo_recipaux_S) The following is provable:
∀x, ∀g : setset, ∀n, nat_p nSNo_recipaux x g (ordsucc n) = (SNo_recipaux x g n 0 SNo_recipauxset (SNo_recipaux x g n 0) x (SNoR x) g SNo_recipauxset (SNo_recipaux x g n 1) x (SNoL_pos x) g,SNo_recipaux x g n 1 SNo_recipauxset (SNo_recipaux x g n 0) x (SNoL_pos x) g SNo_recipauxset (SNo_recipaux x g n 1) x (SNoR x) g)
Proof:
Let x, g and n be given.
Assume Hn.
An exact proof term for the current goal is nat_primrec_S ({0},0) (λk p ⇒ (p 0 SNo_recipauxset (p 0) x (SNoR x) g SNo_recipauxset (p 1) x (SNoL_pos x) g,p 1 SNo_recipauxset (p 0) x (SNoL_pos x) g SNo_recipauxset (p 1) x (SNoR x) g)) n Hn.
Theorem. (SNo_recipaux_lem1) The following is provable:
∀x, SNo x0 < x∀g : setset, (x'SNoS_ (SNoLev x), 0 < x'SNo (g x') x' * g x' = 1)∀k, nat_p k(ySNo_recipaux x g k 0, SNo y x * y < 1) (ySNo_recipaux x g k 1, SNo y 1 < x * y)
Proof:
Let x be given.
Assume Hx Hxpos.
Let g be given.
Assume Hg.
Set r to be the term SNo_recipaux x g.
We prove the intermediate claim L1: x'SNoS_ (SNoLev x), 0 < x'∀y y', SNo yy' = (1 + (x' + - x) * y) * g x'SNo y'.
Let x' be given.
Assume Hx' Hx'pos.
Let y and y' be given.
Assume Hy Hy'.
Apply SNoS_E2 (SNoLev x) (SNoLev_ordinal x Hx) x' Hx' to the current goal.
Assume _ _ Hx'1 _.
rewrite the current goal using Hy' (from left to right).
Apply SNo_mul_SNo to the current goal.
Apply SNo_add_SNo to the current goal.
An exact proof term for the current goal is SNo_1.
Apply SNo_mul_SNo to the current goal.
Apply SNo_add_SNo to the current goal.
An exact proof term for the current goal is Hx'1.
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Hy.
Apply Hg x' Hx' Hx'pos to the current goal.
Assume H _.
An exact proof term for the current goal is H.
We prove the intermediate claim L2: x'SNoS_ (SNoLev x), 0 < x'∀y y', SNo yy' = (1 + (x' + - x) * y) * g x'1 + - x * y' = (1 + - x * y) * (x' + - x) * g x'.
Let x' be given.
Assume Hx' Hx'pos.
Let y and y' be given.
Assume Hy Hy'.
Apply SNoS_E2 (SNoLev x) (SNoLev_ordinal x Hx) x' Hx' to the current goal.
Assume _ _ Hx'1 _.
We will prove 1 + - x * y' = (1 + - x * y) * (x' + - x) * g x'.
Apply Hg x' Hx' Hx'pos to the current goal.
Assume Hgx'1 Hgx'2.
rewrite the current goal using mul_SNo_distrR x' (- x) (g x') Hx'1 (SNo_minus_SNo x Hx) Hgx'1 (from left to right).
We will prove 1 + - x * y' = (1 + - x * y) * (x' * g x' + (- x) * g x').
rewrite the current goal using Hgx'2 (from left to right).
We will prove 1 + - x * y' = (1 + - x * y) * (1 + (- x) * g x').
rewrite the current goal using SNo_foil 1 (- x * y) 1 ((- x) * g x') SNo_1 (SNo_minus_SNo (x * y) (SNo_mul_SNo x y Hx Hy)) SNo_1 (SNo_mul_SNo (- x) (g x') (SNo_minus_SNo x Hx) Hgx'1) (from left to right).
We will prove 1 + - x * y' = 1 * 1 + 1 * (- x) * g x' + (- x * y) * 1 + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_oneL 1 SNo_1 (from left to right).
rewrite the current goal using mul_SNo_oneL ((- x) * g x') (SNo_mul_SNo (- x) (g x') (SNo_minus_SNo x Hx) Hgx'1) (from left to right).
rewrite the current goal using mul_SNo_oneR (- x * y) (SNo_minus_SNo (x * y) (SNo_mul_SNo x y Hx Hy)) (from left to right).
We will prove 1 + - x * y' = 1 + (- x) * g x' + - x * y + (- x * y) * (- x) * g x'.
Use f_equal.
We will prove - x * y' = (- x) * g x' + - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using Hy' (from left to right).
We will prove - x * ((1 + (x' + - x) * y) * g x') = (- x) * g x' + - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_distrR 1 ((x' + - x) * y) (g x') SNo_1 (SNo_mul_SNo (x' + - x) y (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hy) Hgx'1 (from left to right).
We will prove - x * (1 * g x' + ((x' + - x) * y) * g x') = (- x) * g x' + - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_oneL (g x') Hgx'1 (from left to right).
We will prove - x * (g x' + ((x' + - x) * y) * g x') = (- x) * g x' + - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_minus_distrL x (g x' + ((x' + - x) * y) * g x') Hx (SNo_add_SNo (g x') (((x' + - x) * y) * g x') Hgx'1 (SNo_mul_SNo ((x' + - x) * y) (g x') (SNo_mul_SNo (x' + - x) y (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hy) Hgx'1)) (from right to left).
We will prove (- x) * (g x' + ((x' + - x) * y) * g x') = (- x) * g x' + - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_distrL (- x) (g x') (((x' + - x) * y) * g x') (SNo_minus_SNo x Hx) Hgx'1 (SNo_mul_SNo ((x' + - x) * y) (g x') (SNo_mul_SNo (x' + - x) y (SNo_add_SNo x' (- x) Hx'1 (SNo_minus_SNo x Hx)) Hy) Hgx'1) (from left to right).
We will prove (- x) * g x' + (- x) * (((x' + - x) * y) * g x') = (- x) * g x' + - x * y + (- x * y) * (- x) * g x'.
Use f_equal.
We will prove (- x) * (((x' + - x) * y) * g x') = - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_distrR x' (- x) y Hx'1 (SNo_minus_SNo x Hx) Hy (from left to right).
We will prove (- x) * ((x' * y + (- x) * y) * g x') = - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_distrR (x' * y) ((- x) * y) (g x') (SNo_mul_SNo x' y Hx'1 Hy) (SNo_mul_SNo (- x) y (SNo_minus_SNo x Hx) Hy) Hgx'1 (from left to right).
We will prove (- x) * ((x' * y) * g x' + ((- x) * y) * g x') = - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_com x' y Hx'1 Hy (from left to right).
rewrite the current goal using mul_SNo_assoc y x' (g x') Hy Hx'1 Hgx'1 (from right to left).
rewrite the current goal using Hgx'2 (from left to right).
rewrite the current goal using mul_SNo_oneR y Hy (from left to right).
We will prove (- x) * (y + ((- x) * y) * g x') = - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_distrL (- x) y (((- x) * y) * g x') (SNo_minus_SNo x Hx) Hy (SNo_mul_SNo ((- x) * y) (g x') (SNo_mul_SNo (- x) y (SNo_minus_SNo x Hx) Hy) Hgx'1) (from left to right).
We will prove (- x) * y + (- x) * ((- x) * y) * g x' = - x * y + (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_minus_distrL x y Hx Hy (from left to right) at position 1.
We will prove - x * y + (- x) * ((- x) * y) * g x' = - x * y + (- x * y) * (- x) * g x'.
Use f_equal.
We will prove (- x) * ((- x) * y) * g x' = (- x * y) * (- x) * g x'.
rewrite the current goal using mul_SNo_assoc (- x) ((- x) * y) (g x') (SNo_minus_SNo x Hx) (SNo_mul_SNo (- x) y (SNo_minus_SNo x Hx) Hy) Hgx'1 (from left to right).
rewrite the current goal using mul_SNo_assoc (- x * y) (- x) (g x') (SNo_minus_SNo (x * y) (SNo_mul_SNo x y Hx Hy)) (SNo_minus_SNo x Hx) Hgx'1 (from left to right).
Use f_equal.
We will prove (- x) * ((- x) * y) = (- x * y) * (- x).
rewrite the current goal using mul_SNo_com (- x) y (SNo_minus_SNo x Hx) Hy (from left to right).
We will prove (- x) * (y * (- x)) = (- x * y) * (- x).
rewrite the current goal using mul_SNo_minus_distrL x y Hx Hy (from right to left).
We will prove (- x) * (y * (- x)) = ((- x) * y) * (- x).
An exact proof term for the current goal is mul_SNo_assoc (- x) y (- x) (SNo_minus_SNo x Hx) Hy (SNo_minus_SNo x Hx).
Apply nat_ind to the current goal.
Apply andI to the current goal.
Let y be given.
We will prove y r 0 0SNo y x * y < 1.
rewrite the current goal using SNo_recipaux_0 (from left to right).
rewrite the current goal using tuple_2_0_eq (from left to right).
We will prove y {0}SNo y x * y < 1.
Assume H1: y {0}.
rewrite the current goal using SingE 0 y H1 (from left to right).
Apply andI to the current goal.
An exact proof term for the current goal is SNo_0.
We will prove x * 0 < 1.
rewrite the current goal using mul_SNo_zeroR x Hx (from left to right).
An exact proof term for the current goal is SNoLt_0_1.
Let y be given.
We will prove y r 0 1SNo y 1 < x * y.
rewrite the current goal using SNo_recipaux_0 (from left to right).
We will prove y ({0},0) 1SNo y 1 < x * y.
rewrite the current goal using tuple_2_1_eq (from left to right).
We will prove y 0SNo y 1 < x * y.
Assume H1: y 0.
We will prove False.
An exact proof term for the current goal is EmptyE y H1.
Let k be given.
Assume Hk: nat_p k.
Assume IH.
Apply IH to the current goal.
Assume IH1: yr k 0, SNo y x * y < 1.
Assume IH2: yr k 1, SNo y 1 < x * y.
Apply andI to the current goal.
Let y' be given.
We will prove y' r (ordsucc k) 0SNo y' x * y' < 1.
rewrite the current goal using SNo_recipaux_S (from left to right).
We will prove y' (r k 0 SNo_recipauxset (r k 0) x (SNoR x) g SNo_recipauxset (r k 1) x (SNoL_pos x) g,r k 1 SNo_recipauxset (r k 0) x (SNoL_pos x) g SNo_recipauxset (r k 1) x (SNoR x) g) 0SNo y' x * y' < 1.
rewrite the current goal using tuple_2_0_eq (from left to right).
We will prove y' r k 0 SNo_recipauxset (r k 0) x (SNoR x) g SNo_recipauxset (r k 1) x (SNoL_pos x) gSNo y' x * y' < 1.
Assume H1.
Apply binunionE (r k 0 SNo_recipauxset (r k 0) x (SNoR x) g) (SNo_recipauxset (r k 1) x (SNoL_pos x) g) y' H1 to the current goal.
Assume H1.
Apply binunionE (r k 0) (SNo_recipauxset (r k 0) x (SNoR x) g) y' H1 to the current goal.
An exact proof term for the current goal is IH1 y'.
Assume H1.
Apply SNo_recipauxset_E (r k 0) x (SNoR x) g y' H1 to the current goal.
Let y be given.
Assume Hy: y r k 0.
Let x' be given.
Assume Hx': x' SNoR x.
Assume H2: y' = (1 + (x' + - x) * y) * g x'.
Apply SNoR_E x Hx x' Hx' to the current goal.
Assume Hx'1 Hx'2 Hx'3.
We prove the intermediate claim Lx': x' SNoS_ (SNoLev x).
Apply SNoS_I2 x' x Hx'1 Hx Hx'2 to the current goal.
We prove the intermediate claim Lx'pos: 0 < x'.
An exact proof term for the current goal is SNoLt_tra 0 x x' SNo_0 Hx Hx'1 Hxpos Hx'3.
Apply IH1 y Hy to the current goal.
Assume Hy1: SNo y.
Assume Hxy1: x * y < 1.
Apply Hg x' Lx' Lx'pos to the current goal.
Assume Hgx'1: SNo (g x').
Assume Hgx'2: x' * g x' = 1.
We prove the intermediate claim Ly': SNo y'.
An exact proof term for the current goal is L1 x' Lx' Lx'pos y y' Hy1 H2.
Apply andI to the current goal.
An exact proof term for the current goal is Ly'.
Apply SNo_recip_lem3 x x' (g x') y y' Hx Hxpos Hx' Hgx'1 Hgx'2 Hy1 Hxy1 Ly' to the current goal.
We will prove 1 + - x * y' = (1 + - x * y) * (x' + - x) * g x'.
An exact proof term for the current goal is L2 x' Lx' Lx'pos y y' Hy1 H2.
Assume H1.
Apply SNo_recipauxset_E (r k 1) x (SNoL_pos x) g y' H1 to the current goal.
Let y be given.
Assume Hy: y r k 1.
Let x' be given.
Assume Hx': x' SNoL_pos x.
Assume H2: y' = (1 + (x' + - x) * y) * g x'.
Apply SepE (SNoL x) (λw ⇒ 0 < w) x' Hx' to the current goal.
Assume Hx'0: x' SNoL x.
Assume Hx'pos: 0 < x'.
Apply SNoL_E x Hx x' Hx'0 to the current goal.
Assume Hx'1 Hx'2 Hx'3.
We prove the intermediate claim Lx': x' SNoS_ (SNoLev x).
Apply SNoS_I2 x' x Hx'1 Hx Hx'2 to the current goal.
Apply IH2 y Hy to the current goal.
Assume Hy1: SNo y.
Assume Hxy1: 1 < x * y.
Apply Hg x' Lx' Hx'pos to the current goal.
Assume Hgx'1: SNo (g x').
Assume Hgx'2: x' * g x' = 1.
We prove the intermediate claim Ly': SNo y'.
An exact proof term for the current goal is L1 x' Lx' Hx'pos y y' Hy1 H2.
Apply andI to the current goal.
An exact proof term for the current goal is Ly'.
Apply SNo_recip_lem2 x x' (g x') y y' Hx Hxpos Hx' Hgx'1 Hgx'2 Hy1 Hxy1 Ly' to the current goal.
We will prove 1 + - x * y' = (1 + - x * y) * (x' + - x) * g x'.
An exact proof term for the current goal is L2 x' Lx' Hx'pos y y' Hy1 H2.
An exact proof term for the current goal is Hk.
Let y' be given.
We will prove y' r (ordsucc k) 1SNo y' 1 < x * y'.
rewrite the current goal using SNo_recipaux_S (from left to right).
We will prove y' (r k 0 SNo_recipauxset (r k 0) x (SNoR x) g SNo_recipauxset (r k 1) x (SNoL_pos x) g,r k 1 SNo_recipauxset (r k 0) x (SNoL_pos x) g SNo_recipauxset (r k 1) x (SNoR x) g) 1SNo y' 1 < x * y'.
rewrite the current goal using tuple_2_1_eq (from left to right).
We will prove y' r k 1 SNo_recipauxset (r k 0) x (SNoL_pos x) g SNo_recipauxset (r k 1) x (SNoR x) gSNo y' 1 < x * y'.
Assume H1.
Apply binunionE (r k 1 SNo_recipauxset (r k 0) x (SNoL_pos x) g) (SNo_recipauxset (r k 1) x (SNoR x) g) y' H1 to the current goal.
Assume H1.
Apply binunionE (r k 1) (SNo_recipauxset (r k 0) x (SNoL_pos x) g) y' H1 to the current goal.
An exact proof term for the current goal is IH2 y'.
Assume H1.
Apply SNo_recipauxset_E (r k 0) x (SNoL_pos x) g y' H1 to the current goal.
Let y be given.
Assume Hy: y r k 0.
Let x' be given.
Assume Hx': x' SNoL_pos x.
Assume H2: y' = (1 + (x' + - x) * y) * g x'.
Apply SepE (SNoL x) (λw ⇒ 0 < w) x' Hx' to the current goal.
Assume Hx'0: x' SNoL x.
Assume Hx'pos: 0 < x'.
Apply SNoL_E x Hx x' Hx'0 to the current goal.
Assume Hx'1 Hx'2 Hx'3.
We prove the intermediate claim Lx': x' SNoS_ (SNoLev x).
Apply SNoS_I2 x' x Hx'1 Hx Hx'2 to the current goal.
Apply IH1 y Hy to the current goal.
Assume Hy1: SNo y.
Assume Hxy1: x * y < 1.
Apply Hg x' Lx' Hx'pos to the current goal.
Assume Hgx'1: SNo (g x').
Assume Hgx'2: x' * g x' = 1.
We prove the intermediate claim Ly': SNo y'.
An exact proof term for the current goal is L1 x' Lx' Hx'pos y y' Hy1 H2.
Apply andI to the current goal.
An exact proof term for the current goal is Ly'.
Apply SNo_recip_lem1 x x' (g x') y y' Hx Hxpos Hx' Hgx'1 Hgx'2 Hy1 Hxy1 Ly' to the current goal.
We will prove 1 + - x * y' = (1 + - x * y) * (x' + - x) * g x'.
An exact proof term for the current goal is L2 x' Lx' Hx'pos y y' Hy1 H2.
Assume H1.
Apply SNo_recipauxset_E (r k 1) x (SNoR x) g y' H1 to the current goal.
Let y be given.
Assume Hy: y r k 1.
Let x' be given.
Assume Hx': x' SNoR x.
Assume H2: y' = (1 + (x' + - x) * y) * g x'.
Apply SNoR_E x Hx x' Hx' to the current goal.
Assume Hx'1 Hx'2 Hx'3.
We prove the intermediate claim Lx': x' SNoS_ (SNoLev x).
Apply SNoS_I2 x' x Hx'1 Hx Hx'2 to the current goal.
We prove the intermediate claim Lx'pos: 0 < x'.
An exact proof term for the current goal is SNoLt_tra 0 x x' SNo_0 Hx Hx'1 Hxpos Hx'3.
Apply IH2 y Hy to the current goal.
Assume Hy1: SNo y.
Assume Hxy1: 1 < x * y.
Apply Hg x' Lx' Lx'pos to the current goal.
Assume Hgx'1: SNo (g x').
Assume Hgx'2: x' * g x' = 1.
We prove the intermediate claim Ly': SNo y'.
An exact proof term for the current goal is L1 x' Lx' Lx'pos y y' Hy1 H2.
Apply andI to the current goal.
An exact proof term for the current goal is Ly'.
Apply SNo_recip_lem4 x x' (g x') y y' Hx Hxpos Hx' Hgx'1 Hgx'2 Hy1 Hxy1 Ly' to the current goal.
We will prove 1 + - x * y' = (1 + - x * y) * (x' + - x) * g x'.
An exact proof term for the current goal is L2 x' Lx' Lx'pos y y' Hy1 H2.
An exact proof term for the current goal is Hk.
Theorem. (SNo_recipaux_lem2) The following is provable:
∀x, SNo x0 < x∀g : setset, (x'SNoS_ (SNoLev x), 0 < x'SNo (g x') x' * g x' = 1)SNoCutP (kωSNo_recipaux x g k 0) (kωSNo_recipaux x g k 1)
Proof:
Let x be given.
Assume Hx Hxpos.
Let g be given.
Assume Hg.
Set L to be the term kωSNo_recipaux x g k 0.
Set R to be the term kωSNo_recipaux x g k 1.
We will prove (xL, SNo x) (yR, SNo y) (xL, yR, x < y).
We prove the intermediate claim L1: ∀k, nat_p k(ySNo_recipaux x g k 0, SNo y x * y < 1) (ySNo_recipaux x g k 1, SNo y 1 < x * y).
An exact proof term for the current goal is SNo_recipaux_lem1 x Hx Hxpos g Hg.
Apply and3I to the current goal.
Let y be given.
Assume Hy.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x g k 0) y Hy to the current goal.
Let k be given.
Assume Hk.
Assume H1: y SNo_recipaux x g k 0.
Apply L1 k (omega_nat_p k Hk) to the current goal.
Assume H2 _.
Apply H2 y H1 to the current goal.
Assume H3 _.
An exact proof term for the current goal is H3.
Let y be given.
Assume Hy.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x g k 1) y Hy to the current goal.
Let k be given.
Assume Hk.
Assume H1: y SNo_recipaux x g k 1.
Apply L1 k (omega_nat_p k Hk) to the current goal.
Assume _ H2.
Apply H2 y H1 to the current goal.
Assume H3 _.
An exact proof term for the current goal is H3.
Let w be given.
Assume Hw.
Let z be given.
Assume Hz.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x g k 0) w Hw to the current goal.
Let k be given.
Assume Hk.
Assume H1: w SNo_recipaux x g k 0.
Apply L1 k (omega_nat_p k Hk) to the current goal.
Assume H2 _.
Apply H2 w H1 to the current goal.
Assume H3: SNo w.
Assume H4: x * w < 1.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x g k 1) z Hz to the current goal.
Let k' be given.
Assume Hk'.
Assume H5: z SNo_recipaux x g k' 1.
Apply L1 k' (omega_nat_p k' Hk') to the current goal.
Assume _ H6.
Apply H6 z H5 to the current goal.
Assume H7: SNo z.
Assume H8: 1 < x * z.
We will prove w < z.
Apply SNoLtLe_or w z H3 H7 to the current goal.
Assume H9: w < z.
An exact proof term for the current goal is H9.
Assume H9: z w.
We will prove False.
Apply SNoLt_irref 1 to the current goal.
We will prove 1 < 1.
Apply SNoLt_tra 1 (x * z) 1 SNo_1 (SNo_mul_SNo x z Hx H7) SNo_1 H8 to the current goal.
We will prove x * z < 1.
Apply SNoLeLt_tra (x * z) (x * w) 1 (SNo_mul_SNo x z Hx H7) (SNo_mul_SNo x w Hx H3) SNo_1 to the current goal.
We will prove x * z x * w.
Apply nonneg_mul_SNo_Le x z w Hx (SNoLtLe 0 x Hxpos) H7 H3 H9 to the current goal.
We will prove x * w < 1.
An exact proof term for the current goal is H4.
Theorem. (SNo_recipaux_ext) The following is provable:
∀x, SNo x∀g h : setset, (x'SNoS_ (SNoLev x), g x' = h x')∀k, nat_p kSNo_recipaux x g k = SNo_recipaux x h k
Proof:
Let x be given.
Assume Hx.
Let g and h be given.
Assume Hgh.
Apply nat_ind to the current goal.
We will prove SNo_recipaux x g 0 = SNo_recipaux x h 0.
rewrite the current goal using SNo_recipaux_0 x h (from left to right).
An exact proof term for the current goal is SNo_recipaux_0 x g.
Let k be given.
Assume Hk: nat_p k.
Assume IH: SNo_recipaux x g k = SNo_recipaux x h k.
We will prove SNo_recipaux x g (ordsucc k) = SNo_recipaux x h (ordsucc k).
rewrite the current goal using SNo_recipaux_S x g k Hk (from left to right).
rewrite the current goal using SNo_recipaux_S x h k Hk (from left to right).
rewrite the current goal using IH (from left to right).
We prove the intermediate claim L1: SNo_recipaux x h k 0 SNo_recipauxset (SNo_recipaux x h k 0) x (SNoR x) g SNo_recipauxset (SNo_recipaux x h k 1) x (SNoL_pos x) g = SNo_recipaux x h k 0 SNo_recipauxset (SNo_recipaux x h k 0) x (SNoR x) h SNo_recipauxset (SNo_recipaux x h k 1) x (SNoL_pos x) h.
Use f_equal.
Use f_equal.
We will prove SNo_recipauxset (SNo_recipaux x h k 0) x (SNoR x) g = SNo_recipauxset (SNo_recipaux x h k 0) x (SNoR x) h.
Apply SNo_recipauxset_ext to the current goal.
Let w be given.
Assume Hw: w SNoR x.
We will prove g w = h w.
Apply SNoR_E x Hx w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply Hgh to the current goal.
We will prove w SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 w x Hw1 Hx Hw2.
We will prove SNo_recipauxset (SNo_recipaux x h k 1) x (SNoL_pos x) g = SNo_recipauxset (SNo_recipaux x h k 1) x (SNoL_pos x) h.
Apply SNo_recipauxset_ext to the current goal.
Let w be given.
Assume Hw: w SNoL_pos x.
We will prove g w = h w.
Apply SNoL_E x Hx w (SepE1 (SNoL x) (λw ⇒ 0 < w) w Hw) to the current goal.
Assume Hw1 Hw2 Hw3.
Apply Hgh to the current goal.
We will prove w SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 w x Hw1 Hx Hw2.
We prove the intermediate claim L2: SNo_recipaux x h k 1 SNo_recipauxset (SNo_recipaux x h k 0) x (SNoL_pos x) g SNo_recipauxset (SNo_recipaux x h k 1) x (SNoR x) g = SNo_recipaux x h k 1 SNo_recipauxset (SNo_recipaux x h k 0) x (SNoL_pos x) h SNo_recipauxset (SNo_recipaux x h k 1) x (SNoR x) h.
Use f_equal.
Use f_equal.
We will prove SNo_recipauxset (SNo_recipaux x h k 0) x (SNoL_pos x) g = SNo_recipauxset (SNo_recipaux x h k 0) x (SNoL_pos x) h.
Apply SNo_recipauxset_ext to the current goal.
Let w be given.
Assume Hw: w SNoL_pos x.
We will prove g w = h w.
Apply SNoL_E x Hx w (SepE1 (SNoL x) (λw ⇒ 0 < w) w Hw) to the current goal.
Assume Hw1 Hw2 Hw3.
Apply Hgh to the current goal.
We will prove w SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 w x Hw1 Hx Hw2.
We will prove SNo_recipauxset (SNo_recipaux x h k 1) x (SNoR x) g = SNo_recipauxset (SNo_recipaux x h k 1) x (SNoR x) h.
Apply SNo_recipauxset_ext to the current goal.
Let w be given.
Assume Hw: w SNoR x.
We will prove g w = h w.
Apply SNoR_E x Hx w Hw to the current goal.
Assume Hw1 Hw2 Hw3.
Apply Hgh to the current goal.
We will prove w SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 w x Hw1 Hx Hw2.
rewrite the current goal using L1 (from left to right).
rewrite the current goal using L2 (from left to right).
Use reflexivity.
Beginning of Section recip_SNo_pos
Let G : set(setset)setλx g ⇒ SNoCut (kωSNo_recipaux x g k 0) (kωSNo_recipaux x g k 1)
Definition. We define recip_SNo_pos to be SNo_rec_i G of type setset.
Theorem. (recip_SNo_pos_eq) The following is provable:
∀x, SNo xrecip_SNo_pos x = G x recip_SNo_pos
Proof:
Apply SNo_rec_i_eq G to the current goal.
Let x be given.
Assume Hx.
Let g and h be given.
Assume Hgh.
Use f_equal.
Apply famunion_ext to the current goal.
Let k be given.
Assume Hk.
We will prove SNo_recipaux x g k 0 = SNo_recipaux x h k 0.
Use f_equal.
Apply SNo_recipaux_ext x Hx g h to the current goal.
We will prove wSNoS_ (SNoLev x), g w = h w.
An exact proof term for the current goal is Hgh.
An exact proof term for the current goal is omega_nat_p k Hk.
Apply famunion_ext to the current goal.
Let k be given.
Assume Hk.
We will prove SNo_recipaux x g k 1 = SNo_recipaux x h k 1.
Use f_equal.
Apply SNo_recipaux_ext x Hx g h to the current goal.
We will prove wSNoS_ (SNoLev x), g w = h w.
An exact proof term for the current goal is Hgh.
An exact proof term for the current goal is omega_nat_p k Hk.
Theorem. (recip_SNo_pos_prop1) The following is provable:
∀x, SNo x0 < xSNo (recip_SNo_pos x) x * recip_SNo_pos x = 1
Proof:
Apply SNoLev_ind to the current goal.
Let x be given.
Assume Hx: SNo x.
Assume IH: wSNoS_ (SNoLev x), 0 < wSNo (recip_SNo_pos w) w * recip_SNo_pos w = 1.
Assume Hxpos: 0 < x.
We will prove SNo (recip_SNo_pos x) x * recip_SNo_pos x = 1.
rewrite the current goal using recip_SNo_pos_eq x Hx (from left to right).
We will prove SNo (G x recip_SNo_pos) x * G x recip_SNo_pos = 1.
Set L to be the term kωSNo_recipaux x recip_SNo_pos k 0.
Set R to be the term kωSNo_recipaux x recip_SNo_pos k 1.
We prove the intermediate claim L1: ∀k, nat_p k(ySNo_recipaux x recip_SNo_pos k 0, SNo y x * y < 1) (ySNo_recipaux x recip_SNo_pos k 1, SNo y 1 < x * y).
An exact proof term for the current goal is SNo_recipaux_lem1 x Hx Hxpos recip_SNo_pos IH.
We prove the intermediate claim L1L: wL, x * w < 1.
Let w be given.
Assume Hw.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 0) w Hw to the current goal.
Let k be given.
Assume Hk: k ω.
Assume H1: w SNo_recipaux x recip_SNo_pos k 0.
Apply L1 k (omega_nat_p k Hk) to the current goal.
Assume H2 _.
Apply H2 w H1 to the current goal.
Assume _ H3.
An exact proof term for the current goal is H3.
We prove the intermediate claim L1R: zR, 1 < x * z.
Let z be given.
Assume Hz.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 1) z Hz to the current goal.
Let k be given.
Assume Hk: k ω.
Assume H1: z SNo_recipaux x recip_SNo_pos k 1.
Apply L1 k (omega_nat_p k Hk) to the current goal.
Assume _ H2.
Apply H2 z H1 to the current goal.
Assume _ H3.
An exact proof term for the current goal is H3.
We prove the intermediate claim L2: SNoCutP L R.
An exact proof term for the current goal is SNo_recipaux_lem2 x Hx Hxpos recip_SNo_pos IH.
Apply L2 to the current goal.
Assume HLHR.
Apply HLHR to the current goal.
Assume HL: wL, SNo w.
Assume HR: zR, SNo z.
Assume HLR: wL, zR, w < z.
Set y to be the term SNoCut L R.
Apply SNoCutP_SNoCut_impred L R L2 to the current goal.
Assume H1: SNo y.
Assume H3: wL, w < y.
Assume H4: zR, y < z.
Assume H5: ∀u, SNo u(wL, w < u)(zR, u < z)SNoLev y SNoLev u SNoEq_ (SNoLev y) y u.
We prove the intermediate claim L3: SNo (x * y).
An exact proof term for the current goal is SNo_mul_SNo x y Hx H1.
We prove the intermediate claim L4: 0 < y.
Apply H3 to the current goal.
We will prove 0 L.
Apply famunionI ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 0) 0 to the current goal.
We will prove 0 ω.
Apply nat_p_omega to the current goal.
An exact proof term for the current goal is nat_0.
We will prove 0 SNo_recipaux x recip_SNo_pos 0 0.
rewrite the current goal using SNo_recipaux_0 (from left to right).
We will prove 0 ({0},0) 0.
rewrite the current goal using tuple_2_0_eq (from left to right).
Apply SingI to the current goal.
We prove the intermediate claim L5: 0 < x * y.
An exact proof term for the current goal is mul_SNo_pos_pos x y Hx H1 Hxpos L4.
We prove the intermediate claim L6: wSNoL y, w'L, w w'.
An exact proof term for the current goal is SNoL_SNoCutP_ex L R L2.
We prove the intermediate claim L7: zSNoR y, z'R, z' z.
An exact proof term for the current goal is SNoR_SNoCutP_ex L R L2.
Apply andI to the current goal.
We will prove SNo y.
An exact proof term for the current goal is H1.
We will prove x * y = 1.
Apply dneg to the current goal.
Assume HC: x * y 1.
Apply SNoLt_trichotomy_or_impred (x * y) 1 L3 SNo_1 to the current goal.
Assume H6: x * y < 1.
We prove the intermediate claim L8: 1 SNoR (x * y).
Apply SNoR_I to the current goal.
An exact proof term for the current goal is L3.
An exact proof term for the current goal is SNo_1.
We will prove SNoLev 1 SNoLev (x * y).
rewrite the current goal using ordinal_SNoLev 1 ordinal_1 (from left to right).
We will prove 1 SNoLev (x * y).
Apply ordinal_In_Or_Subq 1 (SNoLev (x * y)) ordinal_1 (SNoLev_ordinal (x * y) L3) to the current goal.
Assume H7: 1 SNoLev (x * y).
An exact proof term for the current goal is H7.
Assume H7: SNoLev (x * y) 1.
We will prove False.
Apply HC to the current goal.
We will prove x * y = 1.
An exact proof term for the current goal is pos_low_eq_one (x * y) L3 L5 H7.
An exact proof term for the current goal is H6.
We prove the intermediate claim L9: ∀v w w', SNo vSNo wSNo w'v SNoS_ (SNoLev x)0 < vv * y + x * w 1 + v * w(1 + (v + - x) * w') * recip_SNo_pos v L(- v + x) * w' (- v + x) * wFalse.
Let v, w and w' be given.
Assume Hv1 Hw1 Hw' HvS Hvpos H7 Hw'' H8.
Set w'' to be the term (1 + (v + - x) * w') * recip_SNo_pos v.
We prove the intermediate claim Lw''1: SNo w''.
An exact proof term for the current goal is HL w'' Hw''.
Apply SNoLt_irref 1 to the current goal.
We will prove 1 < 1.
Apply SNoLtLe_tra 1 (1 + v * (y + - w'')) 1 SNo_1 (SNo_add_SNo 1 (v * (y + - w'')) SNo_1 (SNo_mul_SNo v (y + - w'') Hv1 (SNo_add_SNo y (- w'') H1 (SNo_minus_SNo w'' Lw''1)))) SNo_1 to the current goal.
We will prove 1 < 1 + v * (y + - w'').
rewrite the current goal using add_SNo_0R 1 SNo_1 (from right to left) at position 1.
We will prove 1 + 0 < 1 + v * (y + - w'').
Apply add_SNo_Lt2 1 0 (v * (y + - w'')) SNo_1 SNo_0 (SNo_mul_SNo v (y + - w'') Hv1 (SNo_add_SNo y (- w'') H1 (SNo_minus_SNo w'' Lw''1))) to the current goal.
We will prove 0 < v * (y + - w'').
Apply mul_SNo_pos_pos v (y + - w'') Hv1 (SNo_add_SNo y (- w'') H1 (SNo_minus_SNo w'' Lw''1)) Hvpos to the current goal.
We will prove 0 < y + - w''.
Apply SNoLt_minus_pos w'' y Lw''1 H1 to the current goal.
We will prove w'' < y.
An exact proof term for the current goal is H3 w'' Hw''.
We will prove 1 + v * (y + - w'') 1.
rewrite the current goal using mul_SNo_distrL v y (- w'') Hv1 H1 (SNo_minus_SNo w'' Lw''1) (from left to right).
We will prove 1 + v * y + v * (- w'') 1.
rewrite the current goal using mul_SNo_minus_distrR v w'' Hv1 Lw''1 (from left to right).
We will prove 1 + v * y + - v * w'' 1.
We will prove 1 + v * y + - v * (1 + (v + - x) * w') * recip_SNo_pos v 1.
Apply IH v HvS Hvpos to the current goal.
Assume Hrv1: SNo (recip_SNo_pos v).
Assume Hrv2: v * recip_SNo_pos v = 1.
rewrite the current goal using mul_SNo_com (1 + (v + - x) * w') (recip_SNo_pos v) (SNo_add_SNo 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) Hrv1 (from left to right).
We will prove 1 + v * y + - v * recip_SNo_pos v * (1 + (v + - x) * w') 1.
rewrite the current goal using mul_SNo_assoc v (recip_SNo_pos v) (1 + (v + - x) * w') Hv1 Hrv1 (SNo_add_SNo 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (from left to right).
We will prove 1 + v * y + - (v * recip_SNo_pos v) * (1 + (v + - x) * w') 1.
rewrite the current goal using Hrv2 (from left to right).
We will prove 1 + v * y + - 1 * (1 + (v + - x) * w') 1.
rewrite the current goal using mul_SNo_oneL (1 + (v + - x) * w') (SNo_add_SNo 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (from left to right).
We will prove 1 + v * y + - (1 + (v + - x) * w') 1.
rewrite the current goal using minus_add_SNo_distr 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw') (from left to right).
We will prove 1 + v * y + - 1 + - (v + - x) * w' 1.
rewrite the current goal using add_SNo_rotate_3_1 (- 1) (- (v + - x) * w') (v * y) (SNo_minus_SNo 1 SNo_1) (SNo_minus_SNo ((v + - x) * w') (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (SNo_mul_SNo v y Hv1 H1) (from right to left).
We will prove 1 + - 1 + - (v + - x) * w' + v * y 1.
rewrite the current goal using add_SNo_minus_SNo_prop2 1 (- (v + - x) * w' + v * y) SNo_1 (SNo_add_SNo (- (v + - x) * w') (v * y) (SNo_minus_SNo ((v + - x) * w') (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (SNo_mul_SNo v y Hv1 H1)) (from left to right).
We will prove - (v + - x) * w' + v * y 1.
rewrite the current goal using mul_SNo_minus_distrL (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw' (from right to left).
We will prove (- (v + - x)) * w' + v * y 1.
rewrite the current goal using minus_add_SNo_distr v (- x) Hv1 (SNo_minus_SNo x Hx) (from left to right).
rewrite the current goal using minus_SNo_invol x Hx (from left to right).
We will prove (- v + x) * w' + v * y 1.
Apply SNoLe_tra ((- v + x) * w' + v * y) ((- v + x) * w + v * y) 1 (SNo_add_SNo ((- v + x) * w') (v * y) (SNo_mul_SNo (- v + x) w' (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw') (SNo_mul_SNo v y Hv1 H1)) (SNo_add_SNo ((- v + x) * w) (v * y) (SNo_mul_SNo (- v + x) w (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw1) (SNo_mul_SNo v y Hv1 H1)) SNo_1 to the current goal.
We will prove (- v + x) * w' + v * y (- v + x) * w + v * y.
Apply add_SNo_Le1 ((- v + x) * w') (v * y) ((- v + x) * w) (SNo_mul_SNo (- v + x) w' (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw') (SNo_mul_SNo v y Hv1 H1) (SNo_mul_SNo (- v + x) w (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw1) to the current goal.
We will prove (- v + x) * w' (- v + x) * w.
An exact proof term for the current goal is H8.
We will prove (- v + x) * w + v * y 1.
rewrite the current goal using mul_SNo_distrR (- v) x w (SNo_minus_SNo v Hv1) Hx Hw1 (from left to right).
We will prove ((- v) * w + x * w) + v * y 1.
rewrite the current goal using mul_SNo_minus_distrL v w Hv1 Hw1 (from left to right).
We will prove (- v * w + x * w) + v * y 1.
rewrite the current goal using add_SNo_assoc (- v * w) (x * w) (v * y) (SNo_minus_SNo (v * w) (SNo_mul_SNo v w Hv1 Hw1)) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1) (from right to left).
We will prove - v * w + x * w + v * y 1.
rewrite the current goal using add_SNo_com (- v * w) (x * w + v * y) (SNo_minus_SNo (v * w) (SNo_mul_SNo v w Hv1 Hw1)) (SNo_add_SNo (x * w) (v * y) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1)) (from left to right).
Apply add_SNo_minus_Le2 1 (- v * w) (x * w + v * y) SNo_1 (SNo_minus_SNo (v * w) (SNo_mul_SNo v w Hv1 Hw1)) (SNo_add_SNo (x * w) (v * y) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1)) to the current goal.
We will prove x * w + v * y 1 + - - v * w.
rewrite the current goal using minus_SNo_invol (v * w) (SNo_mul_SNo v w Hv1 Hw1) (from left to right).
We will prove x * w + v * y 1 + v * w.
rewrite the current goal using add_SNo_com (x * w) (v * y) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1) (from left to right).
We will prove v * y + x * w 1 + v * w.
An exact proof term for the current goal is H7.
Apply mul_SNo_SNoR_interpolate_impred x y Hx H1 1 L8 to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoR y.
Assume H7: v * y + x * w 1 + v * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1: SNo v.
Assume Hv2: SNoLev v SNoLev x.
Assume Hv3: v < x.
Apply SNoR_E y H1 w Hw to the current goal.
Assume Hw1: SNo w.
Assume Hw2: SNoLev w SNoLev y.
Assume Hw3: y < w.
We prove the intermediate claim LvS: v SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 v x Hv1 Hx Hv2.
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.
Apply L7 w Hw to the current goal.
Let w' be given.
Assume Hw'.
Apply Hw' to the current goal.
Assume Hw'1: w' R.
Assume Hw'2: w' w.
We prove the intermediate claim Lw': SNo w'.
An exact proof term for the current goal is HR w' Hw'1.
We prove the intermediate claim Lxw': SNo (x * w').
An exact proof term for the current goal is SNo_mul_SNo x w' Hx Lw'.
We prove the intermediate claim Lvpos: 0 < v.
Apply SNoLtLe_or 0 v SNo_0 Hv1 to the current goal.
Assume H8: 0 < v.
An exact proof term for the current goal is H8.
Assume H8: v 0.
We will prove False.
Apply SNoLt_irref 1 to the current goal.
We will prove 1 < 1.
Apply SNoLtLe_tra 1 (x * w') 1 SNo_1 Lxw' SNo_1 to the current goal.
We will prove 1 < x * w'.
An exact proof term for the current goal is L1R w' Hw'1.
We will prove x * w' 1.
Apply SNoLe_tra (x * w') (x * w) 1 Lxw' Lxw SNo_1 to the current goal.
We will prove x * w' x * w.
An exact proof term for the current goal is nonneg_mul_SNo_Le x w' w Hx (SNoLtLe 0 x Hxpos) Lw' Hw1 Hw'2.
We will prove x * w 1.
Apply SNoLe_tra (x * w) (v * (y + - w) + x * w) 1 Lxw (SNo_add_SNo (v * (y + - w)) (x * w) (SNo_mul_SNo v (y + - w) Hv1 (SNo_add_SNo y (- w) H1 (SNo_minus_SNo w Hw1))) Lxw) SNo_1 to the current goal.
We will prove x * w v * (y + - w) + x * w.
rewrite the current goal using add_SNo_0L (x * w) Lxw (from right to left) at position 1.
We will prove 0 + x * w v * (y + - w) + x * w.
Apply add_SNo_Le1 0 (x * w) (v * (y + - w)) SNo_0 Lxw (SNo_mul_SNo v (y + - w) Hv1 (SNo_add_SNo y (- w) H1 (SNo_minus_SNo w Hw1))) to the current goal.
We will prove 0 v * (y + - w).
Apply mul_SNo_nonpos_neg v (y + - w) Hv1 (SNo_add_SNo y (- w) H1 (SNo_minus_SNo w Hw1)) H8 to the current goal.
We will prove y + - w < 0.
Apply add_SNo_minus_Lt1b y w 0 H1 Hw1 SNo_0 to the current goal.
We will prove y < 0 + w.
rewrite the current goal using add_SNo_0L w Hw1 (from left to right).
We will prove y < w.
An exact proof term for the current goal is Hw3.
We will prove v * (y + - w) + x * w 1.
rewrite the current goal using mul_SNo_distrL v y (- w) Hv1 H1 (SNo_minus_SNo w Hw1) (from left to right).
We will prove (v * y + v * (- w)) + x * w 1.
rewrite the current goal using add_SNo_com_3b_1_2 (v * y) (v * (- w)) (x * w) (SNo_mul_SNo v y Hv1 H1) (SNo_mul_SNo v (- w) Hv1 (SNo_minus_SNo w Hw1)) Lxw (from left to right).
We will prove (v * y + x * w) + v * (- w) 1.
Apply add_SNo_minus_Le2 1 (v * (- w)) (v * y + x * w) SNo_1 (SNo_mul_SNo v (- w) Hv1 (SNo_minus_SNo w Hw1)) (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 H1) Lxw) to the current goal.
We will prove v * y + x * w 1 + - v * (- w).
rewrite the current goal using mul_SNo_minus_distrR v w Hv1 Hw1 (from left to right).
rewrite the current goal using minus_SNo_invol (v * w) (SNo_mul_SNo v w Hv1 Hw1) (from left to right).
An exact proof term for the current goal is H7.
Set w'' to be the term (1 + (v + - x) * w') * recip_SNo_pos v.
We prove the intermediate claim Lw'': w'' L.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 1) w' Hw'1 to the current goal.
Let k be given.
Assume Hk: k ω.
Assume H9: w' SNo_recipaux x recip_SNo_pos k 1.
Apply famunionI ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 0) (ordsucc k) w'' (omega_ordsucc k Hk) to the current goal.
We will prove w'' SNo_recipaux x recip_SNo_pos (ordsucc k) 0.
rewrite the current goal using SNo_recipaux_S x recip_SNo_pos k (omega_nat_p k Hk) (from left to right).
rewrite the current goal using tuple_2_0_eq (from left to right).
Apply binunionI2 to the current goal.
Apply SNo_recipauxset_I (SNo_recipaux x recip_SNo_pos k 1) x (SNoL_pos x) recip_SNo_pos to the current goal.
We will prove w' SNo_recipaux x recip_SNo_pos k 1.
An exact proof term for the current goal is H9.
We will prove v SNoL_pos x.
We will prove v {wSNoL x|0 < w}.
Apply SepI to the current goal.
An exact proof term for the current goal is Hv.
We will prove 0 < v.
An exact proof term for the current goal is Lvpos.
Apply L9 v w w' Hv1 Hw1 Lw' LvS Lvpos H7 Lw'' to the current goal.
We will prove (- v + x) * w' (- v + x) * w.
Apply nonneg_mul_SNo_Le (- v + x) w' w (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) to the current goal.
We will prove 0 - v + x.
rewrite the current goal using add_SNo_com (- v) x (SNo_minus_SNo v Hv1) Hx (from left to right).
We will prove 0 x + - v.
Apply add_SNo_minus_Le2b x v 0 Hx Hv1 SNo_0 to the current goal.
We will prove 0 + v x.
rewrite the current goal using add_SNo_0L v Hv1 (from left to right).
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 Lw'.
An exact proof term for the current goal is Hw1.
An exact proof term for the current goal is Hw'2.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoL y.
Assume H7: v * y + x * w 1 + v * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1: SNo v.
Assume Hv2: SNoLev v SNoLev x.
Assume Hv3: x < v.
Apply SNoL_E y H1 w Hw to the current goal.
Assume Hw1: SNo w.
Assume Hw2: SNoLev w SNoLev y.
Assume Hw3: w < y.
We prove the intermediate claim LvS: v SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 v x Hv1 Hx Hv2.
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.
Apply L6 w Hw to the current goal.
Let w' be given.
Assume Hw'.
Apply Hw' to the current goal.
Assume Hw'1: w' L.
Assume Hw'2: w w'.
We prove the intermediate claim Lw': SNo w'.
An exact proof term for the current goal is HL w' Hw'1.
We prove the intermediate claim Lxw': SNo (x * w').
An exact proof term for the current goal is SNo_mul_SNo x w' Hx Lw'.
We prove the intermediate claim Lvpos: 0 < v.
An exact proof term for the current goal is SNoLt_tra 0 x v SNo_0 Hx Hv1 Hxpos Hv3.
Set w'' to be the term (1 + (v + - x) * w') * recip_SNo_pos v.
We prove the intermediate claim Lw'': w'' L.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 0) w' Hw'1 to the current goal.
Let k be given.
Assume Hk: k ω.
Assume H9: w' SNo_recipaux x recip_SNo_pos k 0.
Apply famunionI ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 0) (ordsucc k) w'' (omega_ordsucc k Hk) to the current goal.
We will prove w'' SNo_recipaux x recip_SNo_pos (ordsucc k) 0.
rewrite the current goal using SNo_recipaux_S x recip_SNo_pos k (omega_nat_p k Hk) (from left to right).
rewrite the current goal using tuple_2_0_eq (from left to right).
Apply binunionI1 to the current goal.
Apply binunionI2 to the current goal.
Apply SNo_recipauxset_I (SNo_recipaux x recip_SNo_pos k 0) x (SNoR x) recip_SNo_pos to the current goal.
We will prove w' SNo_recipaux x recip_SNo_pos k 0.
An exact proof term for the current goal is H9.
We will prove v SNoR x.
An exact proof term for the current goal is Hv.
Apply L9 v w w' Hv1 Hw1 Lw' LvS Lvpos H7 Lw'' to the current goal.
We will prove (- v + x) * w' (- v + x) * w.
Apply nonpos_mul_SNo_Le (- v + x) w' w (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) to the current goal.
We will prove - v + x 0.
rewrite the current goal using add_SNo_com (- v) x (SNo_minus_SNo v Hv1) Hx (from left to right).
We will prove x + - v 0.
Apply add_SNo_minus_Le2 0 (- v) x SNo_0 (SNo_minus_SNo v Hv1) Hx to the current goal.
rewrite the current goal using minus_SNo_invol v Hv1 (from left to right).
We will prove x 0 + v.
rewrite the current goal using add_SNo_0L v Hv1 (from left to right).
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 Lw'.
An exact proof term for the current goal is Hw1.
An exact proof term for the current goal is Hw'2.
Assume H6: x * y = 1.
Apply HC to the current goal.
An exact proof term for the current goal is H6.
Assume H6: 1 < x * y.
We prove the intermediate claim L10: 1 SNoL (x * y).
Apply SNoL_I to the current goal.
An exact proof term for the current goal is L3.
An exact proof term for the current goal is SNo_1.
We will prove SNoLev 1 SNoLev (x * y).
rewrite the current goal using ordinal_SNoLev 1 ordinal_1 (from left to right).
We will prove 1 SNoLev (x * y).
Apply ordinal_In_Or_Subq 1 (SNoLev (x * y)) ordinal_1 (SNoLev_ordinal (x * y) L3) to the current goal.
Assume H7: 1 SNoLev (x * y).
An exact proof term for the current goal is H7.
Assume H7: SNoLev (x * y) 1.
We will prove False.
Apply HC to the current goal.
We will prove x * y = 1.
An exact proof term for the current goal is pos_low_eq_one (x * y) L3 L5 H7.
An exact proof term for the current goal is H6.
We prove the intermediate claim L11: ∀v w w', SNo vSNo wSNo w'v SNoS_ (SNoLev x)0 < v1 + v * w v * y + x * w(1 + (v + - x) * w') * recip_SNo_pos v R(- v + x) * w (- v + x) * w'False.
Let v, w and w' be given.
Assume Hv1 Hw1 Hw' HvS Hvpos H7 Hw'' H8.
Set w'' to be the term (1 + (v + - x) * w') * recip_SNo_pos v.
We prove the intermediate claim Lw''1: SNo w''.
An exact proof term for the current goal is HR w'' Hw''.
Apply SNoLt_irref 1 to the current goal.
We will prove 1 < 1.
Apply SNoLeLt_tra 1 (1 + v * (y + - w'')) 1 SNo_1 (SNo_add_SNo 1 (v * (y + - w'')) SNo_1 (SNo_mul_SNo v (y + - w'') Hv1 (SNo_add_SNo y (- w'') H1 (SNo_minus_SNo w'' Lw''1)))) SNo_1 to the current goal.
We will prove 1 1 + v * (y + - w'').
rewrite the current goal using mul_SNo_distrL v y (- w'') Hv1 H1 (SNo_minus_SNo w'' Lw''1) (from left to right).
We will prove 1 1 + v * y + v * (- w'').
rewrite the current goal using mul_SNo_minus_distrR v w'' Hv1 Lw''1 (from left to right).
We will prove 1 1 + v * y + - v * w''.
We will prove 1 1 + v * y + - v * (1 + (v + - x) * w') * recip_SNo_pos v.
Apply IH v HvS Hvpos to the current goal.
Assume Hrv1: SNo (recip_SNo_pos v).
Assume Hrv2: v * recip_SNo_pos v = 1.
rewrite the current goal using mul_SNo_com (1 + (v + - x) * w') (recip_SNo_pos v) (SNo_add_SNo 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) Hrv1 (from left to right).
We will prove 1 1 + v * y + - v * recip_SNo_pos v * (1 + (v + - x) * w').
rewrite the current goal using mul_SNo_assoc v (recip_SNo_pos v) (1 + (v + - x) * w') Hv1 Hrv1 (SNo_add_SNo 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (from left to right).
We will prove 1 1 + v * y + - (v * recip_SNo_pos v) * (1 + (v + - x) * w').
rewrite the current goal using Hrv2 (from left to right).
We will prove 1 1 + v * y + - 1 * (1 + (v + - x) * w').
rewrite the current goal using mul_SNo_oneL (1 + (v + - x) * w') (SNo_add_SNo 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (from left to right).
We will prove 1 1 + v * y + - (1 + (v + - x) * w').
rewrite the current goal using minus_add_SNo_distr 1 ((v + - x) * w') SNo_1 (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw') (from left to right).
We will prove 1 1 + v * y + - 1 + - (v + - x) * w'.
rewrite the current goal using add_SNo_rotate_3_1 (- 1) (- (v + - x) * w') (v * y) (SNo_minus_SNo 1 SNo_1) (SNo_minus_SNo ((v + - x) * w') (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (SNo_mul_SNo v y Hv1 H1) (from right to left).
We will prove 1 1 + - 1 + - (v + - x) * w' + v * y.
rewrite the current goal using add_SNo_minus_SNo_prop2 1 (- (v + - x) * w' + v * y) SNo_1 (SNo_add_SNo (- (v + - x) * w') (v * y) (SNo_minus_SNo ((v + - x) * w') (SNo_mul_SNo (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw')) (SNo_mul_SNo v y Hv1 H1)) (from left to right).
We will prove 1 - (v + - x) * w' + v * y.
rewrite the current goal using mul_SNo_minus_distrL (v + - x) w' (SNo_add_SNo v (- x) Hv1 (SNo_minus_SNo x Hx)) Hw' (from right to left).
We will prove 1 (- (v + - x)) * w' + v * y.
rewrite the current goal using minus_add_SNo_distr v (- x) Hv1 (SNo_minus_SNo x Hx) (from left to right).
rewrite the current goal using minus_SNo_invol x Hx (from left to right).
We will prove 1 (- v + x) * w' + v * y.
Apply SNoLe_tra 1 ((- v + x) * w + v * y) ((- v + x) * w' + v * y) SNo_1 (SNo_add_SNo ((- v + x) * w) (v * y) (SNo_mul_SNo (- v + x) w (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw1) (SNo_mul_SNo v y Hv1 H1)) (SNo_add_SNo ((- v + x) * w') (v * y) (SNo_mul_SNo (- v + x) w' (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw') (SNo_mul_SNo v y Hv1 H1)) to the current goal.
We will prove 1 (- v + x) * w + v * y.
rewrite the current goal using mul_SNo_distrR (- v) x w (SNo_minus_SNo v Hv1) Hx Hw1 (from left to right).
We will prove 1 ((- v) * w + x * w) + v * y.
rewrite the current goal using mul_SNo_minus_distrL v w Hv1 Hw1 (from left to right).
We will prove 1 (- v * w + x * w) + v * y.
rewrite the current goal using add_SNo_assoc (- v * w) (x * w) (v * y) (SNo_minus_SNo (v * w) (SNo_mul_SNo v w Hv1 Hw1)) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1) (from right to left).
We will prove 1 - v * w + x * w + v * y.
rewrite the current goal using add_SNo_com (- v * w) (x * w + v * y) (SNo_minus_SNo (v * w) (SNo_mul_SNo v w Hv1 Hw1)) (SNo_add_SNo (x * w) (v * y) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1)) (from left to right).
We will prove 1 (x * w + v * y) + - v * w.
Apply add_SNo_minus_Le2b (x * w + v * y) (v * w) 1 (SNo_add_SNo (x * w) (v * y) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1)) (SNo_mul_SNo v w Hv1 Hw1) SNo_1 to the current goal.
We will prove 1 + v * w x * w + v * y.
rewrite the current goal using add_SNo_com (x * w) (v * y) (SNo_mul_SNo x w Hx Hw1) (SNo_mul_SNo v y Hv1 H1) (from left to right).
We will prove 1 + v * w v * y + x * w.
An exact proof term for the current goal is H7.
We will prove (- v + x) * w + v * y (- v + x) * w' + v * y.
Apply add_SNo_Le1 ((- v + x) * w) (v * y) ((- v + x) * w') (SNo_mul_SNo (- v + x) w (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw1) (SNo_mul_SNo v y Hv1 H1) (SNo_mul_SNo (- v + x) w' (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) Hw') to the current goal.
We will prove (- v + x) * w (- v + x) * w'.
An exact proof term for the current goal is H8.
We will prove 1 + v * (y + - w'') < 1.
rewrite the current goal using add_SNo_0R 1 SNo_1 (from right to left) at position 4.
We will prove 1 + v * (y + - w'') < 1 + 0.
Apply add_SNo_Lt2 1 (v * (y + - w'')) 0 SNo_1 (SNo_mul_SNo v (y + - w'') Hv1 (SNo_add_SNo y (- w'') H1 (SNo_minus_SNo w'' Lw''1))) SNo_0 to the current goal.
We will prove v * (y + - w'') < 0.
Apply mul_SNo_pos_neg v (y + - w'') Hv1 (SNo_add_SNo y (- w'') H1 (SNo_minus_SNo w'' Lw''1)) Hvpos to the current goal.
We will prove y + - w'' < 0.
Apply add_SNo_minus_Lt1b y w'' 0 H1 Lw''1 SNo_0 to the current goal.
We will prove y < 0 + w''.
rewrite the current goal using add_SNo_0L w'' Lw''1 (from left to right).
We will prove y < w''.
An exact proof term for the current goal is H4 w'' Hw''.
Apply mul_SNo_SNoL_interpolate_impred x y Hx H1 1 L10 to the current goal.
Let v be given.
Assume Hv: v SNoL x.
Let w be given.
Assume Hw: w SNoL y.
Assume H7: 1 + v * w v * y + x * w.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1: SNo v.
Assume Hv2: SNoLev v SNoLev x.
Assume Hv3: v < x.
Apply SNoL_E y H1 w Hw to the current goal.
Assume Hw1: SNo w.
Assume Hw2: SNoLev w SNoLev y.
Assume Hw3: w < y.
We prove the intermediate claim LvS: v SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 v x Hv1 Hx Hv2.
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.
Apply L6 w Hw to the current goal.
Let w' be given.
Assume Hw'.
Apply Hw' to the current goal.
Assume Hw'1: w' L.
Assume Hw'2: w w'.
We prove the intermediate claim Lw': SNo w'.
An exact proof term for the current goal is HL w' Hw'1.
We prove the intermediate claim Lxw': SNo (x * w').
An exact proof term for the current goal is SNo_mul_SNo x w' Hx Lw'.
We prove the intermediate claim Lvpos: 0 < v.
Apply SNoLtLe_or 0 v SNo_0 Hv1 to the current goal.
Assume H8: 0 < v.
An exact proof term for the current goal is H8.
Assume H8: v 0.
We will prove False.
Apply SNoLt_irref 1 to the current goal.
We will prove 1 < 1.
Apply SNoLeLt_tra 1 (x * w') 1 SNo_1 Lxw' SNo_1 to the current goal.
We will prove 1 x * w'.
Apply SNoLe_tra 1 (x * w) (x * w') SNo_1 Lxw Lxw' to the current goal.
We will prove 1 x * w.
Apply SNoLe_tra 1 (v * (y + - w) + x * w) (x * w) SNo_1 (SNo_add_SNo (v * (y + - w)) (x * w) (SNo_mul_SNo v (y + - w) Hv1 (SNo_add_SNo y (- w) H1 (SNo_minus_SNo w Hw1))) Lxw) Lxw to the current goal.
We will prove 1 v * (y + - w) + x * w.
rewrite the current goal using mul_SNo_distrL v y (- w) Hv1 H1 (SNo_minus_SNo w Hw1) (from left to right).
We will prove 1 (v * y + v * (- w)) + x * w.
rewrite the current goal using add_SNo_com_3b_1_2 (v * y) (v * (- w)) (x * w) (SNo_mul_SNo v y Hv1 H1) (SNo_mul_SNo v (- w) Hv1 (SNo_minus_SNo w Hw1)) Lxw (from left to right).
We will prove 1 (v * y + x * w) + v * (- w).
rewrite the current goal using mul_SNo_minus_distrR v w Hv1 Hw1 (from left to right).
We will prove 1 (v * y + x * w) + - v * w.
Apply add_SNo_minus_Le2b (v * y + x * w) (v * w) 1 (SNo_add_SNo (v * y) (x * w) (SNo_mul_SNo v y Hv1 H1) Lxw) (SNo_mul_SNo v w Hv1 Hw1) SNo_1 to the current goal.
We will prove 1 + v * w v * y + x * w.
An exact proof term for the current goal is H7.
We will prove v * (y + - w) + x * w x * w.
rewrite the current goal using add_SNo_0L (x * w) Lxw (from right to left) at position 2.
We will prove v * (y + - w) + x * w 0 + x * w.
Apply add_SNo_Le1 (v * (y + - w)) (x * w) 0 (SNo_mul_SNo v (y + - w) Hv1 (SNo_add_SNo y (- w) H1 (SNo_minus_SNo w Hw1))) Lxw SNo_0 to the current goal.
We will prove v * (y + - w) 0.
Apply mul_SNo_nonpos_pos v (y + - w) Hv1 (SNo_add_SNo y (- w) H1 (SNo_minus_SNo w Hw1)) H8 to the current goal.
We will prove 0 < y + - w.
Apply add_SNo_minus_Lt2b y w 0 H1 Hw1 SNo_0 to the current goal.
We will prove 0 + w < y.
rewrite the current goal using add_SNo_0L w Hw1 (from left to right).
We will prove w < y.
An exact proof term for the current goal is Hw3.
We will prove x * w x * w'.
An exact proof term for the current goal is nonneg_mul_SNo_Le x w w' Hx (SNoLtLe 0 x Hxpos) Hw1 Lw' Hw'2.
We will prove x * w' < 1.
An exact proof term for the current goal is L1L w' Hw'1.
Set w'' to be the term (1 + (v + - x) * w') * recip_SNo_pos v.
We prove the intermediate claim Lw'': w'' R.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 0) w' Hw'1 to the current goal.
Let k be given.
Assume Hk: k ω.
Assume H9: w' SNo_recipaux x recip_SNo_pos k 0.
Apply famunionI ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 1) (ordsucc k) w'' (omega_ordsucc k Hk) to the current goal.
We will prove w'' SNo_recipaux x recip_SNo_pos (ordsucc k) 1.
rewrite the current goal using SNo_recipaux_S x recip_SNo_pos k (omega_nat_p k Hk) (from left to right).
rewrite the current goal using tuple_2_1_eq (from left to right).
Apply binunionI1 to the current goal.
Apply binunionI2 to the current goal.
Apply SNo_recipauxset_I (SNo_recipaux x recip_SNo_pos k 0) x (SNoL_pos x) recip_SNo_pos to the current goal.
We will prove w' SNo_recipaux x recip_SNo_pos k 0.
An exact proof term for the current goal is H9.
We will prove v SNoL_pos x.
We will prove v {wSNoL x|0 < w}.
Apply SepI to the current goal.
An exact proof term for the current goal is Hv.
We will prove 0 < v.
An exact proof term for the current goal is Lvpos.
Apply L11 v w w' Hv1 Hw1 Lw' LvS Lvpos H7 Lw'' to the current goal.
We will prove (- v + x) * w (- v + x) * w'.
Apply nonneg_mul_SNo_Le (- v + x) w w' (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) to the current goal.
We will prove 0 - v + x.
rewrite the current goal using add_SNo_com (- v) x (SNo_minus_SNo v Hv1) Hx (from left to right).
We will prove 0 x + - v.
Apply add_SNo_minus_Le2b x v 0 Hx Hv1 SNo_0 to the current goal.
We will prove 0 + v x.
rewrite the current goal using add_SNo_0L v Hv1 (from left to right).
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 Hw1.
An exact proof term for the current goal is Lw'.
An exact proof term for the current goal is Hw'2.
Let v be given.
Assume Hv: v SNoR x.
Let w be given.
Assume Hw: w SNoR y.
Assume H7: 1 + v * w v * y + x * w.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1: SNo v.
Assume Hv2: SNoLev v SNoLev x.
Assume Hv3: x < v.
Apply SNoR_E y H1 w Hw to the current goal.
Assume Hw1: SNo w.
Assume Hw2: SNoLev w SNoLev y.
Assume Hw3: y < w.
We prove the intermediate claim LvS: v SNoS_ (SNoLev x).
An exact proof term for the current goal is SNoS_I2 v x Hv1 Hx Hv2.
We prove the intermediate claim Lvpos: 0 < v.
An exact proof term for the current goal is SNoLt_tra 0 x v SNo_0 Hx Hv1 Hxpos Hv3.
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.
Apply L7 w Hw to the current goal.
Let w' be given.
Assume Hw'.
Apply Hw' to the current goal.
Assume Hw'1: w' R.
Assume Hw'2: w' w.
We prove the intermediate claim Lw': SNo w'.
An exact proof term for the current goal is HR w' Hw'1.
We prove the intermediate claim Lxw': SNo (x * w').
An exact proof term for the current goal is SNo_mul_SNo x w' Hx Lw'.
Set w'' to be the term (1 + (v + - x) * w') * recip_SNo_pos v.
We prove the intermediate claim Lw'': w'' R.
Apply famunionE_impred ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 1) w' Hw'1 to the current goal.
Let k be given.
Assume Hk: k ω.
Assume H9: w' SNo_recipaux x recip_SNo_pos k 1.
Apply famunionI ω (λk ⇒ SNo_recipaux x recip_SNo_pos k 1) (ordsucc k) w'' (omega_ordsucc k Hk) to the current goal.
We will prove w'' SNo_recipaux x recip_SNo_pos (ordsucc k) 1.
rewrite the current goal using SNo_recipaux_S x recip_SNo_pos k (omega_nat_p k Hk) (from left to right).
rewrite the current goal using tuple_2_1_eq (from left to right).
Apply binunionI2 to the current goal.
Apply SNo_recipauxset_I (SNo_recipaux x recip_SNo_pos k 1) x (SNoR x) recip_SNo_pos to the current goal.
We will prove w' SNo_recipaux x recip_SNo_pos k 1.
An exact proof term for the current goal is H9.
We will prove v SNoR x.
An exact proof term for the current goal is Hv.
Apply L11 v w w' Hv1 Hw1 Lw' LvS Lvpos H7 Lw'' to the current goal.
We will prove (- v + x) * w (- v + x) * w'.
Apply L11 v w w' Hv1 Hw1 Lw' LvS Lvpos H7 Lw'' to the current goal.
We will prove (- v + x) * w (- v + x) * w'.
Apply nonpos_mul_SNo_Le (- v + x) w w' (SNo_add_SNo (- v) x (SNo_minus_SNo v Hv1) Hx) to the current goal.
We will prove - v + x 0.
rewrite the current goal using add_SNo_com (- v) x (SNo_minus_SNo v Hv1) Hx (from left to right).
We will prove x + - v 0.
Apply add_SNo_minus_Le2 0 (- v) x SNo_0 (SNo_minus_SNo v Hv1) Hx to the current goal.
rewrite the current goal using minus_SNo_invol v Hv1 (from left to right).
We will prove x 0 + v.
rewrite the current goal using add_SNo_0L v Hv1 (from left to right).
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 Hw1.
An exact proof term for the current goal is Lw'.
An exact proof term for the current goal is Hw'2.
Theorem. (SNo_recip_SNo_pos) The following is provable:
∀x, SNo x0 < xSNo (recip_SNo_pos x)
Proof:
Let x be given.
Assume Hx Hxpos.
Apply recip_SNo_pos_prop1 x Hx Hxpos to the current goal.
Assume H _.
An exact proof term for the current goal is H.
Theorem. (recip_SNo_pos_invR) The following is provable:
∀x, SNo x0 < xx * recip_SNo_pos x = 1
Proof:
Let x be given.
Assume Hx Hxpos.
Apply recip_SNo_pos_prop1 x Hx Hxpos to the current goal.
Assume _ H.
An exact proof term for the current goal is H.
Theorem. (recip_SNo_pos_1) The following is provable:
Proof:
We will prove 1 * recip_SNo_pos 1 = 1 * 1.
rewrite the current goal using mul_SNo_oneL 1 SNo_1 (from left to right).
We will prove 1 * recip_SNo_pos 1 = 1.
An exact proof term for the current goal is recip_SNo_pos_invR 1 SNo_1 SNoLt_0_1.
Theorem. (recip_SNo_pos_is_pos) The following is provable:
∀x, SNo x0 < x0 < recip_SNo_pos x
Proof:
Let x be given.
Assume Hx Hxpos.
We prove the intermediate claim Lrx: SNo (recip_SNo_pos x).
An exact proof term for the current goal is SNo_recip_SNo_pos x Hx Hxpos.
Apply SNoLt_trichotomy_or_impred (recip_SNo_pos x) 0 Lrx SNo_0 to the current goal.
Assume H1: recip_SNo_pos x < 0.
We will prove False.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
Apply SNoLt_tra 0 1 0 SNo_0 SNo_1 SNo_0 SNoLt_0_1 to the current goal.
We will prove 1 < 0.
rewrite the current goal using recip_SNo_pos_invR x Hx Hxpos (from right to left).
We will prove x * recip_SNo_pos x < 0.
An exact proof term for the current goal is mul_SNo_pos_neg x (recip_SNo_pos x) Hx Lrx Hxpos H1.
Assume H1: recip_SNo_pos x = 0.
We will prove False.
Apply neq_1_0 to the current goal.
We will prove 1 = 0.
rewrite the current goal using recip_SNo_pos_invR x Hx Hxpos (from right to left).
We will prove x * recip_SNo_pos x = 0.
rewrite the current goal using H1 (from left to right).
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
Assume H1.
An exact proof term for the current goal is H1.
Theorem. (recip_SNo_pos_invol) The following is provable:
∀x, SNo x0 < xrecip_SNo_pos (recip_SNo_pos x) = x
Proof:
Let x be given.
Assume Hx Hxpos.
We prove the intermediate claim Lrx: SNo (recip_SNo_pos x).
An exact proof term for the current goal is SNo_recip_SNo_pos x Hx Hxpos.
We prove the intermediate claim Lrxpos: 0 < recip_SNo_pos x.
An exact proof term for the current goal is recip_SNo_pos_is_pos x Hx Hxpos.
We prove the intermediate claim Lrrx: SNo (recip_SNo_pos (recip_SNo_pos x)).
An exact proof term for the current goal is SNo_recip_SNo_pos (recip_SNo_pos x) Lrx Lrxpos.
We prove the intermediate claim Lrxn0: recip_SNo_pos x 0.
Assume Hrx0: recip_SNo_pos x = 0.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
rewrite the current goal using Hrx0 (from right to left) at position 2.
An exact proof term for the current goal is Lrxpos.
We will prove recip_SNo_pos (recip_SNo_pos x) = x.
Apply mul_SNo_nonzero_cancel (recip_SNo_pos x) (recip_SNo_pos (recip_SNo_pos x)) x Lrx Lrxn0 Lrrx Hx to the current goal.
rewrite the current goal using mul_SNo_com (recip_SNo_pos x) x Lrx Hx (from left to right).
rewrite the current goal using recip_SNo_pos_invR x Hx Hxpos (from left to right).
An exact proof term for the current goal is recip_SNo_pos_invR (recip_SNo_pos x) Lrx Lrxpos.
Theorem. (recip_SNo_pos_eps_) The following is provable:
∀n, nat_p nrecip_SNo_pos (eps_ n) = 2 ^ n
Proof:
Let n be given.
Assume Hn.
We prove the intermediate claim Len1: SNo (eps_ n).
An exact proof term for the current goal is SNo_eps_ n (nat_p_omega n Hn).
We prove the intermediate claim Len2: 0 < eps_ n.
An exact proof term for the current goal is SNo_eps_pos n (nat_p_omega n Hn).
We prove the intermediate claim Len3: eps_ n 0.
Assume H1: eps_ 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 < eps_ n.
An exact proof term for the current goal is Len2.
Apply mul_SNo_nonzero_cancel (eps_ n) (recip_SNo_pos (eps_ n)) (2 ^ n) Len1 Len3 (SNo_recip_SNo_pos (eps_ n) Len1 Len2) (SNo_exp_SNo_nat 2 SNo_2 n Hn) to the current goal.
We will prove eps_ n * recip_SNo_pos (eps_ n) = eps_ n * 2 ^ n.
rewrite the current goal using mul_SNo_eps_power_2 n Hn (from left to right).
We will prove eps_ n * recip_SNo_pos (eps_ n) = 1.
An exact proof term for the current goal is recip_SNo_pos_invR (eps_ n) Len1 Len2.
Theorem. (recip_SNo_pos_pow_2) The following is provable:
∀n, nat_p nrecip_SNo_pos (2 ^ n) = eps_ n
Proof:
Let n be given.
Assume Hn.
rewrite the current goal using recip_SNo_pos_eps_ n Hn (from right to left).
An exact proof term for the current goal is recip_SNo_pos_invol (eps_ n) (SNo_eps_ n (nat_p_omega n Hn)) (SNo_eps_pos n (nat_p_omega n Hn)).
Theorem. (recip_SNo_pos_2) The following is provable:
Proof:
rewrite the current goal using exp_SNo_nat_1 2 SNo_2 (from right to left).
An exact proof term for the current goal is recip_SNo_pos_pow_2 1 nat_1.
End of Section recip_SNo_pos
Definition. We define recip_SNo to be λx ⇒ if 0 < x then recip_SNo_pos x else if x < 0 then - recip_SNo_pos (- x) else 0 of type setset.
Theorem. (recip_SNo_poscase) The following is provable:
∀x, 0 < xrecip_SNo x = recip_SNo_pos x
Proof:
Let x be given.
Assume Hxpos.
An exact proof term for the current goal is If_i_1 (0 < x) (recip_SNo_pos x) (if x < 0 then - recip_SNo_pos (- x) else 0) Hxpos.
Theorem. (recip_SNo_negcase) The following is provable:
∀x, SNo xx < 0recip_SNo x = - recip_SNo_pos (- x)
Proof:
Let x be given.
Assume Hx Hxneg.
We prove the intermediate claim L1: ¬ (0 < x).
Assume H1.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
An exact proof term for the current goal is SNoLt_tra 0 x 0 SNo_0 Hx SNo_0 H1 Hxneg.
rewrite the current goal using If_i_0 (0 < x) (recip_SNo_pos x) (if x < 0 then - recip_SNo_pos (- x) else 0) L1 (from left to right).
An exact proof term for the current goal is If_i_1 (x < 0) (- recip_SNo_pos (- x)) 0 Hxneg.
Theorem. (recip_SNo_0) The following is provable:
Proof:
rewrite the current goal using If_i_0 (0 < 0) (recip_SNo_pos 0) (if 0 < 0 then - recip_SNo_pos (- 0) else 0) (SNoLt_irref 0) (from left to right).
We will prove (if 0 < 0 then - recip_SNo_pos (- 0) else 0) = 0.
An exact proof term for the current goal is If_i_0 (0 < 0) (- recip_SNo_pos (- 0)) 0 (SNoLt_irref 0).
Theorem. (recip_SNo_1) The following is provable:
Proof:
rewrite the current goal using recip_SNo_poscase 1 SNoLt_0_1 (from left to right).
An exact proof term for the current goal is recip_SNo_pos_1.
Theorem. (SNo_recip_SNo) The following is provable:
∀x, SNo xSNo (recip_SNo 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.
rewrite the current goal using recip_SNo_negcase x Hx H1 (from left to right).
We will prove SNo (- recip_SNo_pos (- x)).
Apply SNo_minus_SNo to the current goal.
We will prove SNo (recip_SNo_pos (- x)).
Apply SNo_recip_SNo_pos to the current goal.
An exact proof term for the current goal is SNo_minus_SNo x Hx.
We will prove 0 < - x.
Apply minus_SNo_Lt_contra2 x 0 Hx SNo_0 to the current goal.
We will prove x < - 0.
rewrite the current goal using minus_SNo_0 (from left to right).
We will prove x < 0.
An exact proof term for the current goal is H1.
Assume H1: x = 0.
rewrite the current goal using H1 (from left to right).
rewrite the current goal using recip_SNo_0 (from left to right).
An exact proof term for the current goal is SNo_0.
Assume H1: 0 < x.
rewrite the current goal using recip_SNo_poscase x H1 (from left to right).
An exact proof term for the current goal is SNo_recip_SNo_pos x Hx H1.
Theorem. (recip_SNo_invR) The following is provable:
∀x, SNo xx 0x * recip_SNo x = 1
Proof:
Let x be given.
Assume Hx Hx0.
Apply SNoLt_trichotomy_or_impred x 0 Hx SNo_0 to the current goal.
Assume H1: x < 0.
rewrite the current goal using recip_SNo_negcase x Hx H1 (from left to right).
We prove the intermediate claim L1: 0 < - x.
Apply minus_SNo_Lt_contra2 x 0 Hx SNo_0 to the current goal.
We will prove x < - 0.
rewrite the current goal using minus_SNo_0 (from left to right).
We will prove x < 0.
An exact proof term for the current goal is H1.
We will prove x * (- recip_SNo_pos (- x)) = 1.
rewrite the current goal using mul_SNo_minus_distrR x (recip_SNo_pos (- x)) Hx (SNo_recip_SNo_pos (- x) (SNo_minus_SNo x Hx) L1) (from left to right).
We will prove - (x * recip_SNo_pos (- x)) = 1.
rewrite the current goal using mul_SNo_minus_distrL x (recip_SNo_pos (- x)) Hx (SNo_recip_SNo_pos (- x) (SNo_minus_SNo x Hx) L1) (from right to left).
We will prove (- x) * recip_SNo_pos (- x) = 1.
An exact proof term for the current goal is recip_SNo_pos_invR (- x) (SNo_minus_SNo x Hx) L1.
Assume H1: x = 0.
We will prove False.
An exact proof term for the current goal is Hx0 H1.
Assume H1: 0 < x.
rewrite the current goal using recip_SNo_poscase x H1 (from left to right).
An exact proof term for the current goal is recip_SNo_pos_invR x Hx H1.
Theorem. (recip_SNo_invL) The following is provable:
∀x, SNo xx 0recip_SNo x * x = 1
Proof:
Let x be given.
Assume Hx Hx0.
rewrite the current goal using mul_SNo_com (recip_SNo x) x (SNo_recip_SNo x Hx) Hx (from left to right).
We will prove x * recip_SNo x = 1.
An exact proof term for the current goal is recip_SNo_invR x Hx Hx0.
Theorem. (recip_SNo_eps_) The following is provable:
∀n, nat_p nrecip_SNo (eps_ n) = 2 ^ n
Proof:
Let n be given.
Assume Hn.
rewrite the current goal using recip_SNo_poscase (eps_ n) (SNo_eps_pos n (nat_p_omega n Hn)) (from left to right).
An exact proof term for the current goal is recip_SNo_pos_eps_ n Hn.
Theorem. (recip_SNo_pow_2) The following is provable:
∀n, nat_p nrecip_SNo (2 ^ n) = eps_ n
Proof:
Let n be given.
Assume Hn.
We prove the intermediate claim L1: 0 < 2 ^ n.
An exact proof term for the current goal is exp_SNo_nat_pos 2 SNo_2 SNoLt_0_2 n Hn.
rewrite the current goal using recip_SNo_poscase (2 ^ n) L1 (from left to right).
An exact proof term for the current goal is recip_SNo_pos_pow_2 n Hn.
Theorem. (recip_SNo_2) The following is provable:
Proof:
rewrite the current goal using recip_SNo_poscase 2 SNoLt_0_2 (from left to right).
An exact proof term for the current goal is recip_SNo_pos_2.
Theorem. (recip_SNo_invol) The following is provable:
∀x, SNo xrecip_SNo (recip_SNo 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.
rewrite the current goal using recip_SNo_negcase x Hx H1 (from left to right).
We prove the intermediate claim Lmx: SNo (- x).
An exact proof term for the current goal is SNo_minus_SNo x Hx.
We prove the intermediate claim Lmxpos: 0 < - x.
Apply minus_SNo_Lt_contra2 x 0 Hx SNo_0 to the current goal.
We will prove x < - 0.
rewrite the current goal using minus_SNo_0 (from left to right).
We will prove x < 0.
An exact proof term for the current goal is H1.
We prove the intermediate claim Lrmx: SNo (recip_SNo_pos (- x)).
An exact proof term for the current goal is SNo_recip_SNo_pos (- x) Lmx Lmxpos.
We prove the intermediate claim Lmrmx: SNo (- recip_SNo_pos (- x)).
An exact proof term for the current goal is SNo_minus_SNo (recip_SNo_pos (- x)) Lrmx.
We prove the intermediate claim L1: 0 < recip_SNo_pos (- x).
An exact proof term for the current goal is recip_SNo_pos_is_pos (- x) Lmx Lmxpos.
We prove the intermediate claim L2: - recip_SNo_pos (- x) < 0.
Apply minus_SNo_Lt_contra1 0 (recip_SNo_pos (- x)) SNo_0 Lrmx to the current goal.
We will prove - 0 < recip_SNo_pos (- x).
rewrite the current goal using minus_SNo_0 (from left to right).
An exact proof term for the current goal is L1.
rewrite the current goal using recip_SNo_negcase (- (recip_SNo_pos (- x))) Lmrmx L2 (from left to right).
We will prove - recip_SNo_pos (- (- (recip_SNo_pos (- x)))) = x.
rewrite the current goal using minus_SNo_invol (recip_SNo_pos (- x)) Lrmx (from left to right).
rewrite the current goal using recip_SNo_pos_invol (- x) Lmx Lmxpos (from left to right).
An exact proof term for the current goal is minus_SNo_invol x Hx.
Assume H1: x = 0.
rewrite the current goal using H1 (from left to right).
rewrite the current goal using recip_SNo_0 (from left to right).
An exact proof term for the current goal is recip_SNo_0.
Assume H1: 0 < x.
rewrite the current goal using recip_SNo_poscase x H1 (from left to right).
rewrite the current goal using recip_SNo_poscase (recip_SNo_pos x) (recip_SNo_pos_is_pos x Hx H1) (from left to right).
An exact proof term for the current goal is recip_SNo_pos_invol x Hx H1.
Theorem. (recip_SNo_of_pos_is_pos) The following is provable:
∀x, SNo x0 < x0 < recip_SNo x
Proof:
Let x be given.
Assume Hx Hxpos.
rewrite the current goal using recip_SNo_poscase x Hxpos (from left to right).
An exact proof term for the current goal is recip_SNo_pos_is_pos x Hx Hxpos.
Theorem. (recip_SNo_neg') The following is provable:
∀x, SNo xx < 0recip_SNo x < 0
Proof:
Let x be given.
Assume Hx Hxneg.
We prove the intermediate claim Lmx: SNo (- x).
An exact proof term for the current goal is SNo_minus_SNo x Hx.
We prove the intermediate claim Lmxpos: 0 < - x.
Apply minus_SNo_Lt_contra2 x 0 Hx SNo_0 to the current goal.
We will prove x < - 0.
rewrite the current goal using minus_SNo_0 (from left to right).
An exact proof term for the current goal is Hxneg.
We prove the intermediate claim Lrmx: SNo (recip_SNo_pos (- x)).
Apply SNo_recip_SNo_pos to the current goal.
Apply SNo_minus_SNo to the current goal.
An exact proof term for the current goal is Hx.
An exact proof term for the current goal is Lmxpos.
rewrite the current goal using recip_SNo_negcase x Hx Hxneg (from left to right).
We will prove - recip_SNo_pos (- x) < 0.
Apply minus_SNo_Lt_contra1 0 (recip_SNo_pos (- x)) SNo_0 Lrmx to the current goal.
We will prove - 0 < recip_SNo_pos (- x).
rewrite the current goal using minus_SNo_0 (from left to right).
We will prove 0 < recip_SNo_pos (- x).
Apply recip_SNo_pos_is_pos (- x) Lmx to the current goal.
We will prove 0 < - x.
An exact proof term for the current goal is Lmxpos.
Definition. We define div_SNo to be λx y ⇒ x * recip_SNo y of type setsetset.
Notation. We use :/: as an infix operator with priority 353 and no associativity corresponding to applying term div_SNo.
Theorem. (SNo_div_SNo) The following is provable:
∀x y, SNo xSNo ySNo (x :/: y)
Proof:
Let x and y be given.
Assume Hx Hy.
We will prove SNo (x * recip_SNo y).
An exact proof term for the current goal is SNo_mul_SNo x (recip_SNo y) Hx (SNo_recip_SNo y Hy).
Theorem. (div_SNo_0_num) The following is provable:
∀x, SNo x0 :/: x = 0
Proof:
Let x be given.
Assume Hx.
An exact proof term for the current goal is mul_SNo_zeroL (recip_SNo x) (SNo_recip_SNo x Hx).
Theorem. (div_SNo_0_denum) The following is provable:
∀x, SNo xx :/: 0 = 0
Proof:
Let x be given.
Assume Hx.
We will prove x * recip_SNo 0 = 0.
rewrite the current goal using recip_SNo_0 (from left to right).
An exact proof term for the current goal is mul_SNo_zeroR x Hx.
Theorem. (mul_div_SNo_invL) The following is provable:
∀x y, SNo xSNo yy 0(x :/: y) * y = x
Proof:
Let x and y be given.
Assume Hx Hy Hy0.
We will prove (x * recip_SNo y) * y = x.
rewrite the current goal using mul_SNo_assoc x (recip_SNo y) y Hx (SNo_recip_SNo y Hy) Hy (from right to left).
We will prove x * (recip_SNo y * y) = x.
rewrite the current goal using recip_SNo_invL y Hy Hy0 (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. (mul_div_SNo_invR) The following is provable:
∀x y, SNo xSNo yy 0y * (x :/: y) = x
Proof:
Let x and y be given.
Assume Hx Hy Hy0.
rewrite the current goal using mul_SNo_com y (x :/: y) Hy (SNo_div_SNo x y Hx Hy) (from left to right).
An exact proof term for the current goal is mul_div_SNo_invL x y Hx Hy Hy0.
Theorem. (mul_div_SNo_R) 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.
Apply xm (y = 0) to the current goal.
Assume H1: y = 0.
rewrite the current goal using H1 (from left to right).
rewrite the current goal using div_SNo_0_denum x Hx (from left to right).
rewrite the current goal using div_SNo_0_denum (x * z) (SNo_mul_SNo x z Hx Hz) (from left to right).
We will prove 0 * z = 0.
An exact proof term for the current goal is mul_SNo_zeroL z Hz.
Assume H1: y 0.
Apply mul_SNo_nonzero_cancel y ((x :/: y) * z) ((x * z) :/: y) Hy H1 (SNo_mul_SNo (x :/: y) z (SNo_div_SNo x y Hx Hy) Hz) (SNo_div_SNo (x * z) y (SNo_mul_SNo x z Hx Hz) Hy) to the current goal.
We will prove y * (x :/: y) * z = y * (x * z) :/: y.
rewrite the current goal using mul_div_SNo_invR (x * z) y (SNo_mul_SNo x z Hx Hz) Hy H1 (from left to right).
We will prove y * (x :/: y) * z = x * z.
rewrite the current goal using mul_SNo_assoc y (x :/: y) z Hy (SNo_div_SNo x y Hx Hy) Hz (from left to right).
We will prove (y * (x :/: y)) * z = x * z.
Use f_equal.
We will prove y * (x :/: y) = x.
An exact proof term for the current goal is mul_div_SNo_invR x y Hx Hy H1.
Theorem. (mul_div_SNo_L) The following is provable:
∀x y z, SNo xSNo ySNo zz * (x :/: y) = (z * x) :/: y
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
Use transitivity with (x :/: y) * z, and (x * z) :/: y.
An exact proof term for the current goal is mul_SNo_com z (x :/: y) Hz (SNo_div_SNo x y Hx Hy).
An exact proof term for the current goal is mul_div_SNo_R x y z Hx Hy Hz.
Use f_equal.
An exact proof term for the current goal is mul_SNo_com x z Hx Hz.
Theorem. (div_mul_SNo_invL) The following is provable:
∀x y, SNo xSNo yy 0(x * y) :/: y = x
Proof:
Let x and y be given.
Assume Hx Hy Hy0.
rewrite the current goal using mul_div_SNo_R x y y Hx Hy Hy (from right to left).
We will prove (x :/: y) * y = x.
An exact proof term for the current goal is mul_div_SNo_invL x y Hx Hy Hy0.
Theorem. (div_div_SNo) The following is provable:
∀x y z, SNo xSNo ySNo z(x :/: y) :/: z = x :/: (y * z)
Proof:
Let x, y and z be given.
Assume Hx Hy Hz.
We prove the intermediate claim Lxdy: SNo (x :/: y).
An exact proof term for the current goal is SNo_div_SNo x y Hx Hy.
We prove the intermediate claim Lxdydz: SNo ((x :/: y) :/: z).
An exact proof term for the current goal is SNo_div_SNo (x :/: y) z Lxdy Hz.
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 Lxdyz: SNo (x :/: (y * z)).
An exact proof term for the current goal is SNo_div_SNo x (y * z) Hx Lyz.
Apply xm (y = 0) to the current goal.
Assume Hy0: y = 0.
rewrite the current goal using Hy0 (from left to right).
We will prove (x :/: 0) :/: z = x :/: (0 * z).
rewrite the current goal using mul_SNo_zeroL z Hz (from left to right).
rewrite the current goal using div_SNo_0_denum x Hx (from left to right).
We will prove 0 :/: z = 0.
An exact proof term for the current goal is div_SNo_0_num z Hz.
Assume Hy0: y 0.
Apply xm (z = 0) to the current goal.
Assume Hz0: z = 0.
rewrite the current goal using Hz0 (from left to right).
We will prove (x :/: y) :/: 0 = x :/: (y * 0).
rewrite the current goal using mul_SNo_zeroR y Hy (from left to right).
We will prove (x :/: y) :/: 0 = x :/: 0.
rewrite the current goal using div_SNo_0_denum x Hx (from left to right).
An exact proof term for the current goal is div_SNo_0_denum (x :/: y) Lxdy.
Assume Hz0: z 0.
We prove the intermediate claim Lyz0: y * z 0.
Assume H1: y * z = 0.
Apply Hz0 to the current goal.
We will prove z = 0.
Apply mul_SNo_nonzero_cancel y z 0 Hy Hy0 Hz SNo_0 to the current goal.
We will prove y * z = y * 0.
rewrite the current goal using mul_SNo_zeroR y Hy (from left to right).
An exact proof term for the current goal is H1.
Apply mul_SNo_nonzero_cancel z ((x :/: y) :/: z) (x :/: (y * z)) Hz Hz0 Lxdydz Lxdyz to the current goal.
rewrite the current goal using mul_div_SNo_invR (x :/: y) z Lxdy Hz Hz0 (from left to right).
We will prove x :/: y = z * (x :/: (y * z)).
Apply mul_SNo_nonzero_cancel y (x :/: y) (z * (x :/: (y * z))) Hy Hy0 Lxdy (SNo_mul_SNo z (x :/: (y * z)) Hz Lxdyz) to the current goal.
rewrite the current goal using mul_div_SNo_invR x y Hx Hy Hy0 (from left to right).
We will prove x = y * z * (x :/: (y * z)).
rewrite the current goal using mul_SNo_assoc y z (x :/: (y * z)) Hy Hz Lxdyz (from left to right).
Use symmetry.
An exact proof term for the current goal is mul_div_SNo_invR x (y * z) Hx Lyz Lyz0.
Theorem. (mul_div_SNo_both) 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_div_SNo_L z w (x :/: y) Hz Hw (SNo_div_SNo x y Hx Hy) (from left to right).
We will prove ((x :/: y) * z) :/: w = (x * z) :/: (y * w).
rewrite the current goal using mul_div_SNo_R x y z Hx Hy Hz (from left to right).
We will prove ((x * z) :/: y) :/: w = (x * z) :/: (y * w).
An exact proof term for the current goal is div_div_SNo (x * z) y w (SNo_mul_SNo x z Hx Hz) Hy Hw.
Theorem. (recip_SNo_pos_pos) The following is provable:
∀x, SNo x0 < x0 < recip_SNo_pos x
Proof:
Let x be given.
Assume Hx Hxpos.
Apply SNoLtLe_or 0 (recip_SNo_pos x) SNo_0 (SNo_recip_SNo_pos x Hx Hxpos) to the current goal.
Assume H1: 0 < recip_SNo_pos x.
An exact proof term for the current goal is H1.
Assume H1: recip_SNo_pos x 0.
We will prove False.
Apply SNoLt_irref 0 to the current goal.
We will prove 0 < 0.
Apply SNoLtLe_tra 0 1 0 SNo_0 SNo_1 SNo_0 SNoLt_0_1 to the current goal.
We will prove 1 0.
rewrite the current goal using recip_SNo_pos_invR x Hx Hxpos (from right to left).
We will prove x * recip_SNo_pos x 0.
rewrite the current goal using mul_SNo_zeroR x Hx (from right to left).
We will prove x * recip_SNo_pos x x * 0.
An exact proof term for the current goal is nonneg_mul_SNo_Le x (recip_SNo_pos x) 0 Hx (SNoLtLe 0 x Hxpos) (SNo_recip_SNo_pos x Hx Hxpos) SNo_0 H1.
Theorem. (recip_SNo_of_neg_is_neg) The following is provable:
∀x, SNo xx < 0recip_SNo x < 0
Proof:
Let x be given.
Assume Hx Hxneg.
rewrite the current goal using recip_SNo_negcase x Hx Hxneg (from left to right).
We will prove - recip_SNo_pos (- x) < 0.
We prove the intermediate claim Lmx: SNo (- x).
An exact proof term for the current goal is SNo_minus_SNo x Hx.
We prove the intermediate claim L1: 0 < - x.
Apply minus_SNo_Lt_contra2 x 0 Hx SNo_0 to the current goal.
We will prove x < - 0.
rewrite the current goal using minus_SNo_0 (from left to right).
An exact proof term for the current goal is Hxneg.
Apply minus_SNo_Lt_contra1 0 (recip_SNo_pos (- x)) SNo_0 (SNo_recip_SNo_pos (- x) Lmx L1) to the current goal.
We will prove - 0 < recip_SNo_pos (- x).
rewrite the current goal using minus_SNo_0 (from left to right).
We will prove 0 < recip_SNo_pos (- x).
Apply recip_SNo_pos_pos (- x) (SNo_minus_SNo x Hx) to the current goal.
We will prove 0 < - x.
An exact proof term for the current goal is L1.
Theorem. (div_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 Hxpos Hypos.
We will prove 0 < x * recip_SNo y.
Apply mul_SNo_pos_pos x (recip_SNo y) Hx (SNo_recip_SNo y Hy) Hxpos to the current goal.
We will prove 0 < recip_SNo y.
An exact proof term for the current goal is recip_SNo_of_pos_is_pos y Hy Hypos.
Theorem. (div_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 Hxneg Hyneg.
We will prove 0 < x * recip_SNo y.
Apply mul_SNo_neg_neg x (recip_SNo y) Hx (SNo_recip_SNo y Hy) Hxneg to the current goal.
We will prove recip_SNo y < 0.
An exact proof term for the current goal is recip_SNo_neg' y Hy Hyneg.
Theorem. (div_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 Hxpos Hyneg.
We will prove x * recip_SNo y < 0.
Apply mul_SNo_pos_neg x (recip_SNo y) Hx (SNo_recip_SNo y Hy) Hxpos to the current goal.
We will prove recip_SNo y < 0.
An exact proof term for the current goal is recip_SNo_neg' y Hy Hyneg.
Theorem. (div_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 Hxneg Hypos.
We will prove x * recip_SNo y < 0.
Apply mul_SNo_neg_pos x (recip_SNo y) Hx (SNo_recip_SNo y Hy) Hxneg to the current goal.
We will prove 0 < recip_SNo y.
An exact proof term for the current goal is recip_SNo_of_pos_is_pos y Hy Hypos.
Theorem. (div_SNo_pos_LtL) The following is provable:
∀x y z, SNo xSNo ySNo z0 < yx < z * yx :/: y < z
Proof:
Let x, y and z be given.
Assume Hx Hy Hz Hypos.
Assume H1: x < z * y.
We will prove x :/: y < z.
rewrite the current goal using mul_SNo_oneR z Hz (from right to left).
We will prove x :/: y < z * 1.
We prove the intermediate claim Ly0: y 0.
Assume H2: y = 0.
Apply SNoLt_irref y to the current goal.
We will prove y < y.
rewrite the current goal using H2 (from left to right) at position 1.
An exact proof term for the current goal is Hypos.
rewrite the current goal using recip_SNo_invR y Hy Ly0 (from right to left).
We will prove x * recip_SNo y < z * y * recip_SNo y.
rewrite the current goal using mul_SNo_assoc z y (recip_SNo y) Hz Hy (SNo_recip_SNo y Hy) (from left to right).
We will prove x * recip_SNo y < (z * y) * recip_SNo y.
Apply pos_mul_SNo_Lt' x (z * y) (recip_SNo y) Hx (SNo_mul_SNo z y Hz Hy) (SNo_recip_SNo y Hy) to the current goal.
We will prove 0 < recip_SNo y.
An exact proof term for the current goal is recip_SNo_of_pos_is_pos y Hy Hypos.
We will prove x < z * y.
An exact proof term for the current goal is H1.
Theorem. (div_SNo_pos_LtR) The following is provable:
∀x y z, SNo xSNo ySNo z0 < yz * y < xz < x :/: y
Proof:
Let x, y and z be given.
Assume Hx Hy Hz Hypos.
Assume H1: z * y < x.
We will prove z < x :/: y.
rewrite the current goal using mul_SNo_oneR z Hz (from right to left).
We will prove z * 1 < x :/: y.
We prove the intermediate claim Ly0: y 0.
Assume H2: y = 0.
Apply SNoLt_irref y to the current goal.
We will prove y < y.
rewrite the current goal using H2 (from left to right) at position 1.
An exact proof term for the current goal is Hypos.
rewrite the current goal using recip_SNo_invR y Hy Ly0 (from right to left).
We will prove z * y * recip_SNo y < x * recip_SNo y.
rewrite the current goal using mul_SNo_assoc z y (recip_SNo y) Hz Hy (SNo_recip_SNo y Hy) (from left to right).
We will prove (z * y) * recip_SNo y < x * recip_SNo y.
Apply pos_mul_SNo_Lt' (z * y) x (recip_SNo y) (SNo_mul_SNo z y Hz Hy) Hx (SNo_recip_SNo y Hy) to the current goal.
We will prove 0 < recip_SNo y.
An exact proof term for the current goal is recip_SNo_of_pos_is_pos y Hy Hypos.
We will prove z * y < x.
An exact proof term for the current goal is H1.
Theorem. (div_SNo_pos_LtL') The following is provable:
∀x y z, SNo xSNo ySNo z0 < yx :/: y < zx < z * y
Proof:
Let x, y and z be given.
Assume Hx Hy Hz Hypos.
Assume H1: x :/: y < z.
We will prove x < z * y.
rewrite the current goal using mul_SNo_oneR x Hx (from right to left).
We will prove x * 1 < z * y.
We prove the intermediate claim Ly0: y 0.
Assume H2: y = 0.
Apply SNoLt_irref y to the current goal.
We will prove y < y.
rewrite the current goal using H2 (from left to right) at position 1.
An exact proof term for the current goal is Hypos.
rewrite the current goal using recip_SNo_invL y Hy Ly0 (from right to left).
We will prove x * (recip_SNo y * y) < z * y.
rewrite the current goal using mul_SNo_assoc x (recip_SNo y) y Hx (SNo_recip_SNo y Hy) Hy (from left to right).
We will prove (x * recip_SNo y) * y < z * y.
We will prove (x :/: y) * y < z * y.
An exact proof term for the current goal is pos_mul_SNo_Lt' (x :/: y) z y (SNo_div_SNo x y Hx Hy) Hz Hy Hypos H1.
Theorem. (div_SNo_pos_LtR') The following is provable:
∀x y z, SNo xSNo ySNo z0 < yz < x :/: yz * y < x
Proof:
Let x, y and z be given.
Assume Hx Hy Hz Hypos.
Assume H1: z < x :/: y.
We will prove z * y < x.
rewrite the current goal using mul_SNo_oneR x Hx (from right to left).
We will prove z * y < x * 1.
We prove the intermediate claim Ly0: y 0.
Assume H2: y = 0.
Apply SNoLt_irref y to the current goal.
We will prove y < y.
rewrite the current goal using H2 (from left to right) at position 1.
An exact proof term for the current goal is Hypos.
rewrite the current goal using recip_SNo_invL y Hy Ly0 (from right to left).
We will prove z * y < x * (recip_SNo y * y).
rewrite the current goal using mul_SNo_assoc x (recip_SNo y) y Hx (SNo_recip_SNo y Hy) Hy (from left to right).
We will prove z * y < (x * recip_SNo y) * y.
We will prove z * y < (x :/: y) * y.
An exact proof term for the current goal is pos_mul_SNo_Lt' z (x :/: y) y Hz (SNo_div_SNo x y Hx Hy) Hy Hypos H1.
Theorem. (mul_div_SNo_nonzero_eq) The following is provable:
∀x y z, SNo xSNo ySNo zy 0x = y * zx :/: y = z
Proof:
Let x, y and z be given.
Assume Hx Hy Hz Hy0 H1.
Apply mul_SNo_nonzero_cancel y (x :/: y) z Hy Hy0 (SNo_div_SNo x y Hx Hy) Hz to the current goal.
We will prove y * (x :/: y) = y * z.
rewrite the current goal using mul_div_SNo_invR x y Hx Hy Hy0 (from left to right).
An exact proof term for the current goal is H1.
End of Section SurrealDiv