Primitive . The name
Eps_i is a term of type
(set → prop ) → set .
Axiom. (
Eps_i_ax ) We take the following as an axiom:
∀P : set → prop , ∀x : set , P x → P (Eps_i P )
Definition. We define
True to be
∀p : prop , p → p of type
prop .
Definition. We define
False to be
∀p : prop , p of type
prop .
Definition. We define
not to be
λA : prop ⇒ A → False of type
prop → prop .
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 , (A → B → p ) → p of type
prop → prop → prop .
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 , (A → p ) → (B → p ) → p of type
prop → prop → prop .
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 : prop ⇒ and (A → B ) (B → A ) of type
prop → prop → prop .
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 : A → A → prop , Q x y → Q y x of type
A → A → prop .
Definition. We define
neq to be
λx y : A ⇒ ¬ eq x y of type
A → A → prop .
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 : A → B , (∀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 : A → prop ⇒ ∀P : prop , (∀x : A , Q x → P ) → P of type
(A → prop ) → 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 q → p = q
Primitive . The name
In is a term of type
set → set → prop .
Notation . We use
∈ as an infix operator with priority 500 and no associativity corresponding to applying term
In . Furthermore, we may write
∀ x ∈ A , B to mean
∀ x : set, x ∈ A → B .
Definition. We define
Subq to be
λA B ⇒ ∀ x ∈ A , x ∈ B of type
set → set → prop .
Notation . We use
⊆ as an infix operator with priority 500 and no associativity corresponding to applying term
Subq . Furthermore, we may write
∀ x ⊆ A , B to mean
∀ x : set, x ⊆ A → B .
Axiom. (
set_ext ) We take the following as an axiom:
∀X Y : set , X ⊆ Y → Y ⊆ X → X = Y
Axiom. (
In_ind ) We take the following as an axiom:
∀P : set → prop , (∀X : set , (∀ x ∈ X , 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:
Primitive . The name
⋃ is a term of type
set → set .
Axiom. (
UnionEq ) We take the following as an axiom:
Primitive . The name
𝒫 is a term of type
set → set .
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 → (set → set ) → set .
Notation .
{B | x ∈ A } is notation for
Repl A (λ x . B ).
Axiom. (
ReplEq ) We take the following as an axiom:
∀A : set , ∀F : set → set , ∀y : set , y ∈ { F x | x ∈ A } ↔ ∃ x ∈ A , y = F x
Definition. We define
TransSet to be
λU : set ⇒ ∀ x ∈ U , x ⊆ U of type
set → prop .
Definition. We define
Union_closed to be
λU : set ⇒ ∀X : set , X ∈ U → ⋃ X ∈ U of type
set → prop .
Definition. We define
Power_closed to be
λU : set ⇒ ∀X : set , X ∈ U → 𝒫 X ∈ U of type
set → prop .
Definition. We define
Repl_closed to be
λU : set ⇒ ∀X : set , X ∈ U → ∀F : set → set , (∀x : set , x ∈ X → F x ∈ U ) → { F x | x ∈ X } ∈ U of type
set → prop .
Primitive . The name
UnivOf is a term of type
set → set .
Axiom. (
UnivOf_In ) We take the following as an axiom:
Axiom. (
UnivOf_Min ) We take the following as an axiom:
Axiom. (
FalseE ) We take the following as an axiom:
Axiom. (
TrueI ) We take the following as an axiom:
Axiom. (
andI ) We take the following as an axiom:
∀A B : prop , A → B → A ∧ B
Axiom. (
andEL ) We take the following as an axiom:
Axiom. (
andER ) We take the following as an axiom:
Axiom. (
orIL ) We take the following as an axiom:
Axiom. (
orIR ) We take the following as an axiom:
Beginning of Section PropN
Variable P1 P2 P3 : prop
Axiom. (
and3I ) We take the following as an axiom:
P1 → P2 → P3 → P1 ∧ P2 ∧ P3
Axiom. (
and3E ) We take the following as an axiom:
P1 ∧ P2 ∧ P3 → (∀p : prop , (P1 → P2 → P3 → p ) → p )
Axiom. (
or3I1 ) We take the following as an axiom:
Axiom. (
or3I2 ) We take the following as an axiom:
Axiom. (
or3I3 ) We take the following as an axiom:
Axiom. (
or3E ) We take the following as an axiom:
P1 ∨ P2 ∨ P3 → (∀p : prop , (P1 → p ) → (P2 → p ) → (P3 → p ) → p )
Variable P4 : prop
Axiom. (
and4I ) We take the following as an axiom:
P1 → P2 → P3 → P4 → P1 ∧ P2 ∧ P3 ∧ P4
Variable P5 : prop
Axiom. (
and5I ) We take the following as an axiom:
P1 → P2 → P3 → P4 → P5 → P1 ∧ P2 ∧ P3 ∧ P4 ∧ P5
End of Section PropN
Axiom. (
iffI ) We take the following as an axiom:
∀A B : prop , (A → B ) → (B → A ) → (A ↔ B )
Axiom. (
iffEL ) We take the following as an axiom:
∀A B : prop , (A ↔ B ) → A → B
Axiom. (
iffER ) We take the following as an axiom:
∀A B : prop , (A ↔ B ) → B → A
Axiom. (
iff_refl ) We take the following as an axiom:
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 = y → y = z → x = z
Axiom. (
f_eq_i ) We take the following as an axiom:
∀f : set → set , ∀x y, x = y → f x = f y
Axiom. (
neq_i_sym ) We take the following as an axiom:
Definition. We define
nIn to be
λx X ⇒ ¬ In x X of type
set → set → prop .
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 : set → prop , (∃ x, P x ) → P (Eps_i P )
Axiom. (
pred_ext ) We take the following as an axiom:
∀P Q : set → prop , (∀x, P x ↔ Q x ) → P = Q
Axiom. (
prop_ext_2 ) We take the following as an axiom:
∀p q : prop , (p → q ) → (q → p ) → p = q
Axiom. (
Subq_ref ) We take the following as an axiom:
Axiom. (
Subq_tra ) We take the following as an axiom:
∀X Y Z : set , X ⊆ Y → Y ⊆ Z → X ⊆ Z
Axiom. (
Subq_contra ) We take the following as an axiom:
∀X Y z : set , X ⊆ Y → z ∉ Y → z ∉ X
Axiom. (
EmptyE ) We take the following as an axiom:
Axiom. (
Subq_Empty ) We take the following as an axiom:
Axiom. (
Empty_eq ) We take the following as an axiom:
Axiom. (
UnionI ) We take the following as an axiom:
∀X x Y : set , x ∈ Y → Y ∈ X → x ∈ ⋃ X
Axiom. (
UnionE ) We take the following as an axiom:
∀X x : set , x ∈ ⋃ X → ∃ Y : 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 ∈ Y → Y ∈ X → p ) → p
Axiom. (
PowerI ) We take the following as an axiom:
∀X Y : set , Y ⊆ X → Y ∈ 𝒫 X
Axiom. (
PowerE ) We take the following as an axiom:
∀X Y : set , Y ∈ 𝒫 X → Y ⊆ X
Axiom. (
xm ) We take the following as an axiom:
Axiom. (
dneg ) We take the following as an axiom:
Axiom. (
eq_or_nand ) We take the following as an axiom:
Primitive . The name
exactly1of2 is a term of type
prop → prop → prop .
Axiom. (
exactly1of2_E ) We take the following as an axiom:
∀A B : prop , exactly1of2 A B → ∀p : prop , (A → ¬ B → p ) → (¬ A → B → p ) → p
Axiom. (
ReplI ) We take the following as an axiom:
∀A : set , ∀F : set → set , ∀x : set , x ∈ A → F x ∈ { F x | x ∈ A }
Axiom. (
ReplE ) We take the following as an axiom:
∀A : set , ∀F : set → set , ∀y : set , y ∈ { F x | x ∈ A } → ∃ x ∈ A , y = F x
Axiom. (
ReplE_impred ) We take the following as an axiom:
∀A : set , ∀F : set → set , ∀y : set , y ∈ { F x | x ∈ A } → ∀p : prop , (∀x : set , x ∈ A → y = F x → p ) → p
Axiom. (
ReplE' ) We take the following as an axiom:
∀X, ∀f : set → set , ∀p : set → prop , (∀ x ∈ X , p (f x ) ) → ∀ y ∈ { f x | x ∈ X } , p y
Axiom. (
Repl_Empty ) We take the following as an axiom:
Axiom. (
ReplEq_ext ) We take the following as an axiom:
∀X, ∀F G : set → set , (∀ x ∈ X , F x = G x ) → { F x | x ∈ X } = { G x | x ∈ X }
Axiom. (
Repl_inv_eq ) We take the following as an axiom:
∀P : set → prop , ∀f g : set → set , (∀x, P x → g (f x ) = x ) → ∀X, (∀ x ∈ X , P x ) → { g y | y ∈ { f x | x ∈ X } } = X
Axiom. (
Repl_invol_eq ) We take the following as an axiom:
∀P : set → prop , ∀f : set → set , (∀x, P x → f (f x ) = x ) → ∀X, (∀ x ∈ X , P x ) → { f y | y ∈ { f x | x ∈ X } } = X
Primitive . The name
If_i is a term of type
prop → set → set → set .
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_0 ) We take the following as an axiom:
Axiom. (
If_i_1 ) We take the following as an axiom:
Axiom. (
If_i_or ) We take the following as an axiom:
Primitive . The name
UPair is a term of type
set → set → set .
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:
Axiom. (
UPairI2 ) We take the following as an axiom:
Primitive . The name
Sing is a term of type
set → set .
Notation .
{x } is notation for
Sing x .
Axiom. (
SingI ) We take the following as an axiom:
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
set → set → set .
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 ∈ X → z ∈ X ∪ Y
Axiom. (
binunionI2 ) We take the following as an axiom:
∀X Y z : set , z ∈ Y → z ∈ X ∪ Y
Axiom. (
binunionE ) We take the following as an axiom:
∀X Y z : set , z ∈ X ∪ Y → z ∈ X ∨ z ∈ Y
Axiom. (
binunionE' ) We take the following as an axiom:
∀X Y z, ∀p : prop , (z ∈ X → p ) → (z ∈ Y → p ) → (z ∈ X ∪ Y → p )
Axiom. (
binunion_asso ) We take the following as an axiom:
∀X Y Z : set , X ∪ (Y ∪ Z ) = (X ∪ Y ) ∪ Z
Axiom. (
binunion_com ) We take the following as an axiom:
∀X Y : set , X ∪ Y = Y ∪ X
Definition. We define
SetAdjoin to be
λX y ⇒ X ∪ { y } of type
set → set → set .
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 → (set → set ) → 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 : (set → set ) , ∀x y : set , x ∈ X → y ∈ F x → y ∈ ⋃ x ∈ X F x
Axiom. (
famunionE ) We take the following as an axiom:
∀X : set , ∀F : (set → set ) , ∀y : set , y ∈ (⋃ x ∈ X F x ) → ∃ x ∈ X , y ∈ F x
Axiom. (
famunionE_impred ) We take the following as an axiom:
∀X : set , ∀F : (set → set ) , ∀y : set , y ∈ (⋃ x ∈ X F x ) → ∀p : prop , (∀x, x ∈ X → y ∈ F x → p ) → p
Beginning of Section SepSec
Variable X : set
Variable P : set → prop
Let z : set ≝ Eps_i (λz ⇒ z ∈ X ∧ P z )
Primitive . The name
Sep is a term of type
set .
End of Section SepSec
Notation .
{x ∈ A | B } is notation for
Sep A (λ x . B ).
Axiom. (
SepI ) We take the following as an axiom:
∀X : set , ∀P : (set → prop ) , ∀x : set , x ∈ X → P x → x ∈ { x ∈ X | P x }
Axiom. (
SepE ) We take the following as an axiom:
∀X : set , ∀P : (set → prop ) , ∀x : set , x ∈ { x ∈ X | P x } → x ∈ X ∧ P x
Axiom. (
SepE1 ) We take the following as an axiom:
∀X : set , ∀P : (set → prop ) , ∀x : set , x ∈ { x ∈ X | P x } → x ∈ X
Axiom. (
SepE2 ) We take the following as an axiom:
∀X : set , ∀P : (set → prop ) , ∀x : set , x ∈ { x ∈ X | P x } → P x
Axiom. (
Sep_Subq ) We take the following as an axiom:
∀X : set , ∀P : set → prop , { x ∈ X | P x } ⊆ X
Axiom. (
Sep_In_Power ) We take the following as an axiom:
∀X : set , ∀P : set → prop , { x ∈ X | P x } ∈ 𝒫 X
Primitive . The name
ReplSep is a term of type
set → (set → prop ) → (set → set ) → set .
Notation .
{B | x ∈ A , C } is notation for
ReplSep A (λ x . C ) (λ x . B ).
Axiom. (
ReplSepI ) We take the following as an axiom:
∀X : set , ∀P : set → prop , ∀F : set → set , ∀x : set , x ∈ X → P x → F x ∈ { F x | x ∈ X , P x }
Axiom. (
ReplSepE ) We take the following as an axiom:
∀X : set , ∀P : set → prop , ∀F : set → set , ∀y : set , y ∈ { F x | x ∈ X , P x } → ∃ x : set , x ∈ X ∧ P x ∧ y = F x
Axiom. (
ReplSepE_impred ) We take the following as an axiom:
∀X : set , ∀P : set → prop , ∀F : set → set , ∀y : set , y ∈ { F x | x ∈ X , P x } → ∀p : prop , (∀ x ∈ X , P x → y = F x → p ) → p
Primitive . The name
binintersect is a term of type
set → set → set .
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 ∈ X → z ∈ Y → z ∈ X ∩ Y
Axiom. (
binintersectE ) We take the following as an axiom:
∀X Y z, z ∈ X ∩ Y → z ∈ X ∧ z ∈ Y
Primitive . The name
setminus is a term of type
set → set → set .
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. (
In_irref ) We take the following as an axiom:
Axiom. (
In_no2cycle ) We take the following as an axiom:
Primitive . The name
ordsucc is a term of type
set → set .
Axiom. (
ordsuccI1 ) We take the following as an axiom:
Axiom. (
ordsuccI2 ) We take the following as an axiom:
Axiom. (
ordsuccE ) We take the following as an axiom:
Notation . Natural numbers 0,1,2,... are notation for the terms formed using
Empty as 0 and forming successors with
ordsucc .
Axiom. (
ordsucc_inj ) We take the following as an axiom:
Axiom. (
In_0_1 ) We take the following as an axiom:
Axiom. (
In_0_2 ) We take the following as an axiom:
Axiom. (
In_1_2 ) We take the following as an axiom:
Definition. We define
nat_p to be
λn : set ⇒ ∀p : set → prop , p 0 → (∀x : set , p x → p (ordsucc x ) ) → p n of type
set → prop .
Axiom. (
nat_0 ) We take the following as an axiom:
Axiom. (
nat_ordsucc ) We take the following as an axiom:
Axiom. (
nat_1 ) We take the following as an axiom:
Axiom. (
nat_2 ) We take the following as an axiom:
Axiom. (
nat_ind ) We take the following as an axiom:
Axiom. (
nat_inv ) We take the following as an axiom:
Axiom. (
nat_p_trans ) We take the following as an axiom:
Axiom. (
nat_trans ) We take the following as an axiom:
Axiom. (
cases_1 ) We take the following as an axiom:
∀ i ∈ 1 , ∀p : set → prop , p 0 → p i
Axiom. (
cases_2 ) We take the following as an axiom:
∀ i ∈ 2 , ∀p : set → prop , p 0 → p 1 → p i
Axiom. (
cases_3 ) We take the following as an axiom:
∀ i ∈ 3 , ∀p : set → prop , p 0 → p 1 → p 2 → p i
Axiom. (
neq_0_1 ) We take the following as an axiom:
Axiom. (
neq_1_0 ) We take the following as an axiom:
Axiom. (
neq_0_2 ) We take the following as an axiom:
Axiom. (
neq_2_0 ) We take the following as an axiom:
Axiom. (
neq_1_2 ) We take the following as an axiom:
Axiom. (
ZF_closed_E ) We take the following as an axiom:
Primitive . The name
ω is a term of type
set .
Axiom. (
omega_nat_p ) We take the following as an axiom:
Axiom. (
nat_p_omega ) We take the following as an axiom:
Axiom. (
ordinal_1 ) We take the following as an axiom:
Axiom. (
ordinal_2 ) We take the following as an axiom:
Axiom. (
ordinal_ind ) We take the following as an axiom:
∀p : set → prop , (∀alpha, ordinal alpha → (∀ beta ∈ alpha , p beta ) → p alpha ) → ∀alpha, ordinal alpha → p alpha
Definition. We define
inj to be
λX Y f ⇒ (∀ u ∈ X , f u ∈ Y ) ∧ (∀ u v ∈ X , f u = f v → u = v ) of type
set → set → (set → set ) → prop .
Definition. We define
bij to be
λX Y f ⇒ (∀ u ∈ X , f u ∈ Y ) ∧ (∀ u v ∈ X , f u = f v → u = v ) ∧ (∀ w ∈ Y , ∃ u ∈ X , f u = w ) of type
set → set → (set → set ) → prop .
Axiom. (
bijI ) We take the following as an axiom:
∀X Y, ∀f : set → set , (∀ u ∈ X , f u ∈ Y ) → (∀ u v ∈ X , f u = f v → u = v ) → (∀ w ∈ Y , ∃ u ∈ X , f u = w ) → bij X Y f
Axiom. (
bijE ) We take the following as an axiom:
∀X Y, ∀f : set → set , bij X Y f → ∀p : prop , ((∀ u ∈ X , f u ∈ Y ) → (∀ u v ∈ X , f u = f v → u = v ) → (∀ w ∈ Y , ∃ u ∈ X , f u = w ) → p ) → p
Primitive . The name
inv is a term of type
set → (set → set ) → set → set .
Axiom. (
surj_rinv ) We take the following as an axiom:
Axiom. (
inj_linv ) We take the following as an axiom:
∀X, ∀f : set → set , (∀ u v ∈ X , f u = f v → u = v ) → ∀ x ∈ X , inv X f (f x ) = x
Axiom. (
bij_inv ) We take the following as an axiom:
∀X Y, ∀f : set → set , bij X Y f → bij Y X (inv X f )
Axiom. (
bij_id ) We take the following as an axiom:
Axiom. (
bij_comp ) We take the following as an axiom:
∀X Y Z : set , ∀f g : set → set , bij X Y f → bij Y Z g → bij X Z (λx ⇒ g (f x ) )
Definition. We define
equip to be
λX Y : set ⇒ ∃ f : set → set , bij X Y f of type
set → set → prop .
Axiom. (
equip_ref ) We take the following as an axiom:
Axiom. (
equip_sym ) We take the following as an axiom:
Axiom. (
equip_tra ) We take the following as an axiom:
Beginning of Section SchroederBernstein
End of Section SchroederBernstein
Beginning of Section PigeonHole
End of Section PigeonHole
Definition. We define
finite to be
λX ⇒ ∃ n ∈ ω , equip X n of type
set → prop .
Axiom. (
finite_ind ) We take the following as an axiom:
Axiom. (
Subq_finite ) We take the following as an axiom:
Axiom. (
exandE_i ) We take the following as an axiom:
∀P Q : set → prop , (∃ x, P x ∧ Q x ) → ∀r : prop , (∀x, P x → Q x → r ) → r
Axiom. (
exandE_ii ) We take the following as an axiom:
∀P Q : (set → set ) → prop , (∃ x : set → set , P x ∧ Q x ) → ∀p : prop , (∀x : set → set , P x → Q x → p ) → p
Axiom. (
exandE_iii ) We take the following as an axiom:
∀P Q : (set → set → set ) → prop , (∃ x : set → set → set , P x ∧ Q x ) → ∀p : prop , (∀x : set → set → set , P x → Q x → p ) → p
Axiom. (
exandE_iiii ) We take the following as an axiom:
∀P Q : (set → set → set → set ) → prop , (∃ x : set → set → set → set , P x ∧ Q x ) → ∀p : prop , (∀x : set → set → set → set , P x → Q x → p ) → p
Beginning of Section Descr_ii
Variable P : (set → set ) → prop
Primitive . The name
Descr_ii is a term of type
set → set .
Hypothesis Pex : ∃ f : set → set , P f
Hypothesis Puniq : ∀f g : set → set , P f → P g → f = g
End of Section Descr_ii
Beginning of Section Descr_iii
Variable P : (set → set → set ) → prop
Primitive . The name
Descr_iii is a term of type
set → set → set .
Hypothesis Pex : ∃ f : set → set → set , P f
Hypothesis Puniq : ∀f g : set → set → set , P f → P g → f = g
End of Section Descr_iii
Beginning of Section Descr_Vo1
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 f → P g → f = g
End of Section Descr_Vo1
Beginning of Section If_ii
Variable p : prop
Variable f g : set → set
Primitive . The name
If_ii is a term of type
set → set .
Axiom. (
If_ii_1 ) We take the following as an axiom:
Axiom. (
If_ii_0 ) We take the following as an axiom:
End of Section If_ii
Beginning of Section If_iii
Variable p : prop
Variable f g : set → set → set
Primitive . The name
If_iii is a term of type
set → set → set .
Axiom. (
If_iii_1 ) We take the following as an axiom:
Axiom. (
If_iii_0 ) We take the following as an axiom:
End of Section If_iii
Beginning of Section EpsilonRec_i
Variable F : set → (set → set ) → set
Primitive . The name
In_rec_i is a term of type
set → set .
Hypothesis Fr : ∀X : set , ∀g h : set → set , (∀ x ∈ X , g x = h x ) → F X g = F X h
Axiom. (
In_rec_i_eq ) We take the following as an axiom:
End of Section EpsilonRec_i
Beginning of Section EpsilonRec_ii
Variable F : set → (set → (set → set ) ) → (set → set )
Primitive . The name
In_rec_ii is a term of type
set → (set → set ) .
Hypothesis Fr : ∀X : set , ∀g h : set → (set → set ) , (∀ x ∈ X , g x = h x ) → F X g = F X h
End of Section EpsilonRec_ii
Beginning of Section EpsilonRec_iii
Variable F : set → (set → (set → set → set ) ) → (set → set → set )
Primitive . The name
In_rec_iii is a term of type
set → (set → set → set ) .
Hypothesis Fr : ∀X : set , ∀g h : set → (set → set → set ) , (∀ x ∈ X , g x = h x ) → F X g = F X h
End of Section EpsilonRec_iii
Beginning of Section NatRec
Variable z : set
Variable f : set → set → set
Let F : set → (set → set ) → set ≝ λn g ⇒ if ⋃ n ∈ n then f (⋃ n ) (g (⋃ n ) ) else z
Axiom. (
nat_primrec_r ) We take the following as an axiom:
∀X : set , ∀g h : set → set , (∀ x ∈ X , g x = h x ) → F X g = F X h
End of Section NatRec
Beginning of Section NatArith
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:
Axiom. (
add_nat_SR ) We take the following as an axiom:
Axiom. (
add_nat_p ) We take the following as an axiom:
Definition. We define
mul_nat to be
λn m : set ⇒ nat_primrec 0 (λ_ r ⇒ n + r ) m of type
set → set → set .
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:
Axiom. (
mul_nat_SR ) We take the following as an axiom:
Axiom. (
mul_nat_p ) We take the following as an axiom:
End of Section NatArith
Axiom. (
Inj1_eq ) We take the following as an axiom:
Axiom. (
Inj1I1 ) We take the following as an axiom:
Axiom. (
Inj1I2 ) We take the following as an axiom:
Axiom. (
Inj1E ) We take the following as an axiom:
Axiom. (
Inj1NE1 ) We take the following as an axiom:
Axiom. (
Inj1NE2 ) We take the following as an axiom:
Definition. We define
Inj0 to be
λX ⇒ { Inj1 x | x ∈ X } of type
set → set .
Axiom. (
Inj0I ) We take the following as an axiom:
Axiom. (
Inj0E ) We take the following as an axiom:
Axiom. (
Unj_eq ) We take the following as an axiom:
Axiom. (
Unj_Inj1_eq ) We take the following as an axiom:
Axiom. (
Inj1_inj ) We take the following as an axiom:
Axiom. (
Unj_Inj0_eq ) We take the following as an axiom:
Axiom. (
Inj0_inj ) We take the following as an axiom:
Axiom. (
Inj0_0 ) We take the following as an axiom:
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 ∈ X → Inj0 x ∈ X + Y
Axiom. (
Inj1_setsum ) We take the following as an axiom:
∀X Y y : set , y ∈ Y → Inj1 y ∈ X + Y
Axiom. (
setsum_0_0 ) We take the following as an axiom:
Beginning of Section pair_setsum
Axiom. (
pairI0 ) We take the following as an axiom:
∀X Y x, x ∈ X → pair 0 x ∈ pair X Y
Axiom. (
pairI1 ) We take the following as an axiom:
∀X Y y, y ∈ Y → pair 1 y ∈ pair X Y
Axiom. (
pairE ) We take the following as an axiom:
∀X Y z, z ∈ pair X Y → (∃ x ∈ X , z = pair 0 x ) ∨ (∃ y ∈ Y , z = pair 1 y )
Axiom. (
pairE0 ) We take the following as an axiom:
∀X Y x, pair 0 x ∈ pair X Y → x ∈ X
Axiom. (
pairE1 ) We take the following as an axiom:
∀X Y y, pair 1 y ∈ pair X Y → y ∈ Y
Axiom. (
proj0I ) We take the following as an axiom:
Axiom. (
proj0E ) We take the following as an axiom:
Axiom. (
proj1I ) We take the following as an axiom:
Axiom. (
proj1E ) We take the following as an axiom:
Definition. We define
Sigma to be
λX Y ⇒ ⋃ x ∈ X { pair x y | y ∈ Y x } of type
set → (set → set ) → 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 : set → set , ∀ x ∈ X , ∀ y ∈ Y x , pair x y ∈ ∑ x ∈ X , Y x
Axiom. (
proj0_Sigma ) We take the following as an axiom:
∀X : set , ∀Y : set → set , ∀z : set , z ∈ (∑ x ∈ X , Y x ) → proj0 z ∈ X
Axiom. (
proj1_Sigma ) We take the following as an axiom:
∀X : set , ∀Y : set → set , ∀z : set , z ∈ (∑ x ∈ X , Y x ) → proj1 z ∈ Y (proj0 z )
Axiom. (
pair_Sigma_E1 ) We take the following as an axiom:
∀X : set , ∀Y : set → set , ∀x y : set , pair x y ∈ (∑ x ∈ X , Y x ) → y ∈ Y x
Axiom. (
Sigma_E ) We take the following as an axiom:
∀X : set , ∀Y : set → set , ∀z : set , z ∈ (∑ x ∈ X , Y x ) → ∃ x ∈ X , ∃ y ∈ Y x , z = pair x y
Definition. We define
setprod to be
λX Y : set ⇒ ∑ x ∈ X , Y of type
set → set → set .
Notation . We use
⨯ as an infix operator with priority 440 and which associates to the left corresponding to applying term
setprod .
Let lam : set → (set → set ) → set ≝ Sigma
Definition. We define
ap to be
λf x ⇒ { proj1 z | z ∈ f , ∃ y : set , z = pair x y } of type
set → set → set .
Notation . When
x is a set, a term
x y is notation for
ap x y .
Notation .
λ x ∈ A ⇒ B is notation for the set
Sigma A (λ x : 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 : set → set , ∀ x ∈ X , ∀ y ∈ F x , pair x y ∈ λ x ∈ X ⇒ F x
Axiom. (
lamE ) We take the following as an axiom:
∀X : set , ∀F : set → set , ∀z : set , z ∈ (λ x ∈ X ⇒ F x ) → ∃ x ∈ X , ∃ y ∈ F x , z = pair x y
Axiom. (
apI ) We take the following as an axiom:
∀f x y, pair x y ∈ f → y ∈ f x
Axiom. (
apE ) We take the following as an axiom:
∀f x y, y ∈ f x → pair x y ∈ f
Axiom. (
beta ) We take the following as an axiom:
∀X : set , ∀F : set → set , ∀x : set , x ∈ X → (λ x ∈ X ⇒ F x ) x = F x
Axiom. (
proj0_ap_0 ) We take the following as an axiom:
Axiom. (
proj1_ap_1 ) We take the following as an axiom:
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 : set → set , ∀z : set , z ∈ (∑ x ∈ X , Y x ) → (z 0 ) ∈ X
Axiom. (
ap1_Sigma ) We take the following as an axiom:
∀X : set , ∀Y : set → set , ∀z : set , z ∈ (∑ x ∈ X , Y x ) → (z 1 ) ∈ (Y (z 0 ) )
Definition. We define
pair_p to be
λu : set ⇒ pair (u 0 ) (u 1 ) = u of type
set → prop .
Axiom. (
pair_p_I ) 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 ∈ 𝒫 (∑ x ∈ X , ⋃ (Y x ) ) | ∀ x ∈ X , f x ∈ Y x } of type
set → (set → set ) → 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 : set → set , ∀f : set , (∀ u ∈ f , pair_p u ∧ u 0 ∈ X ) → (∀ x ∈ X , f x ∈ Y x ) → f ∈ ∏ x ∈ X , Y x
Axiom. (
lam_Pi ) We take the following as an axiom:
∀X : set , ∀Y : set → set , ∀F : set → set , (∀ x ∈ X , F x ∈ Y x ) → (λ x ∈ X ⇒ F x ) ∈ (∏ x ∈ X , Y x )
Axiom. (
ap_Pi ) We take the following as an axiom:
∀X : set , ∀Y : set → set , ∀f : set , ∀x : set , f ∈ (∏ x ∈ X , Y x ) → x ∈ X → f x ∈ Y x
Definition. We define
setexp to be
λX Y : set ⇒ ∏ y ∈ Y , X of type
set → set → set .
Notation . We use
:^: as an infix operator with priority 430 and which associates to the left corresponding to applying term
setexp .
Axiom. (
lamI2 ) We take the following as an axiom:
Beginning of Section Tuples
Variable x0 x1 : set
End of Section Tuples
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 → (set → prop ) → prop ) → set .
Primitive . The name
DescrR_i_io_2 is a term of type
(set → (set → prop ) → prop ) → set → prop .
Definition. We define
PNoEq_ to be
λalpha p q ⇒ ∀ beta ∈ alpha , p beta ↔ q beta of type
set → (set → prop ) → (set → prop ) → prop .
Axiom. (
PNoEq_ref_ ) We take the following as an axiom:
∀alpha, ∀p : set → prop , PNoEq_ alpha p p
Axiom. (
PNoEq_sym_ ) We take the following as an axiom:
∀alpha, ∀p q : set → prop , PNoEq_ alpha p q → PNoEq_ alpha q p
Axiom. (
PNoEq_tra_ ) We take the following as an axiom:
Axiom. (
PNoLt_E_ ) We take the following as an axiom:
Axiom. (
PNoLt_irref_ ) We take the following as an axiom:
∀alpha, ∀p : set → prop , ¬ PNoLt_ alpha p p
Axiom. (
PNoLt_mon_ ) We take the following as an axiom:
Axiom. (
PNoLt_tra_ ) We take the following as an axiom:
Primitive . The name
PNoLt is a term of type
set → (set → prop ) → set → (set → prop ) → prop .
Axiom. (
PNoLtI1 ) We take the following as an axiom:
Axiom. (
PNoLtI2 ) We take the following as an axiom:
Axiom. (
PNoLtI3 ) We take the following as an axiom:
Axiom. (
PNoLtE ) We take the following as an axiom:
Axiom. (
PNoLt_irref ) We take the following as an axiom:
∀alpha, ∀p : set → prop , ¬ PNoLt alpha p alpha p
Axiom. (
PNoLtEq_tra ) We take the following as an axiom:
Axiom. (
PNoEqLt_tra ) We take the following as an axiom:
Axiom. (
PNoLt_tra ) We take the following as an axiom:
Definition. We define
PNoLe to be
λalpha p beta q ⇒ PNoLt alpha p beta q ∨ alpha = beta ∧ PNoEq_ alpha p q of type
set → (set → prop ) → set → (set → prop ) → prop .
Axiom. (
PNoLeI1 ) We take the following as an axiom:
Axiom. (
PNoLeI2 ) We take the following as an axiom:
∀alpha, ∀p q : set → prop , PNoEq_ alpha p q → PNoLe alpha p alpha q
Axiom. (
PNoLe_ref ) We take the following as an axiom:
∀alpha, ∀p : set → prop , PNoLe alpha p alpha p
Axiom. (
PNoLtLe_tra ) We take the following as an axiom:
Axiom. (
PNoLeLt_tra ) We take the following as an axiom:
Axiom. (
PNoEqLe_tra ) We take the following as an axiom:
Axiom. (
PNoLe_tra ) We take the following as an axiom:
Axiom. (
PNoLe_downc ) We take the following as an axiom:
Axiom. (
PNo_downc_ref ) We take the following as an axiom:
∀L : set → (set → prop ) → prop , ∀alpha, ordinal alpha → ∀p : set → prop , L alpha p → PNo_downc L alpha p
Axiom. (
PNo_upc_ref ) We take the following as an axiom:
∀R : set → (set → prop ) → prop , ∀alpha, ordinal alpha → ∀p : set → prop , R alpha p → PNo_upc R alpha p
Axiom. (
PNoLe_upc ) We take the following as an axiom:
Definition. We define
PNoLt_pwise to be
λL R ⇒ ∀gamma, ordinal gamma → ∀p : set → prop , L gamma p → ∀delta, ordinal delta → ∀q : set → prop , R delta q → PNoLt gamma p delta q of type
(set → (set → prop ) → prop ) → (set → (set → prop ) → prop ) → prop .
Axiom. (
PNo_extend0_eq ) We take the following as an axiom:
∀alpha, ∀p : set → prop , PNoEq_ alpha p (λdelta ⇒ p delta ∧ delta ≠ alpha )
Axiom. (
PNo_extend1_eq ) We take the following as an axiom:
∀alpha, ∀p : set → prop , PNoEq_ alpha p (λdelta ⇒ p delta ∨ delta = alpha )
Definition. We define
PNo_lenbdd to be
λalpha L ⇒ ∀beta, ∀p : set → prop , L beta p → beta ∈ alpha of type
set → (set → (set → prop ) → prop ) → 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 → (set → prop ) → prop ) → (set → (set → prop ) → prop ) → set → (set → prop ) → prop .
Primitive . The name
PNo_bd is a term of type
(set → (set → prop ) → prop ) → (set → (set → prop ) → prop ) → set .
Primitive . The name
PNo_pred is a term of type
(set → (set → prop ) → prop ) → (set → (set → prop ) → prop ) → set → prop .
Axiom. (
PNo_bd_pred ) We take the following as an axiom:
Axiom. (
PNo_bd_In ) We take the following as an axiom:
Beginning of Section TaggedSets
Notation . We use
' as a postfix operator with priority 100 corresponding to applying term
tag .
Definition. We define
SNoElts_ to be
λalpha ⇒ alpha ∪ { beta ' | beta ∈ alpha } of type
set → set .
Axiom. (
SNoElts_mon ) We take the following as an axiom:
Axiom. (
PNoEq_PSNo ) We take the following as an axiom:
Axiom. (
SNo_PSNo ) We take the following as an axiom:
Primitive . The name
SNo is a term of type
set → prop .
Axiom. (
SNo_SNo ) We take the following as an axiom:
Primitive . The name
SNoLev is a term of type
set → set .
Axiom. (
SNoLev_uniq ) We take the following as an axiom:
Axiom. (
SNoLev_prop ) We take the following as an axiom:
Axiom. (
SNoLev_ ) We take the following as an axiom:
Axiom. (
SNoLev_PSNo ) We take the following as an axiom:
Axiom. (
SNo_Subq ) We take the following as an axiom:
Definition. We define
SNoEq_ to be
λalpha x y ⇒ PNoEq_ alpha (λbeta ⇒ beta ∈ x ) (λbeta ⇒ beta ∈ y ) of type
set → set → set → prop .
Axiom. (
SNoEq_I ) We take the following as an axiom:
Axiom. (
SNo_eq ) We take the following as an axiom:
End of Section TaggedSets
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 .
Axiom. (
SNoLtLe ) We take the following as an axiom:
Axiom. (
SNoLeE ) We take the following as an axiom:
Axiom. (
SNoEq_sym_ ) We take the following as an axiom:
Axiom. (
SNoEq_tra_ ) We take the following as an axiom:
Axiom. (
SNoLtE ) We take the following as an axiom:
Axiom. (
SNoLtI2 ) We take the following as an axiom:
Axiom. (
SNoLtI3 ) We take the following as an axiom:
Axiom. (
SNoLt_irref ) We take the following as an axiom:
Axiom. (
SNoLt_tra ) We take the following as an axiom:
Axiom. (
SNoLe_ref ) We take the following as an axiom:
Axiom. (
SNoLtLe_tra ) We take the following as an axiom:
Axiom. (
SNoLeLt_tra ) We take the following as an axiom:
Axiom. (
SNoLe_tra ) We take the following as an axiom:
Axiom. (
SNoLtLe_or ) We take the following as an axiom:
Axiom. (
SNoCutP_L_0 ) We take the following as an axiom:
Axiom. (
SNoCutP_0_R ) We take the following as an axiom:
Axiom. (
SNoCutP_0_0 ) We take the following as an axiom:
Axiom. (
SNoS_E ) We take the following as an axiom:
Beginning of Section TaggedSets2
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:
Axiom. (
SNoS_I2 ) We take the following as an axiom:
Axiom. (
SNoS_Subq ) We take the following as an axiom:
Axiom. (
SNoS_E2 ) We take the following as an axiom:
Axiom. (
SNoS_In_neq ) We take the following as an axiom:
Axiom. (
SNoS_SNoLev ) We take the following as an axiom:
Axiom. (
SNoL_E ) We take the following as an axiom:
Axiom. (
SNoR_E ) We take the following as an axiom:
Axiom. (
SNoL_SNoS_ ) We take the following as an axiom:
Axiom. (
SNoR_SNoS_ ) We take the following as an axiom:
Axiom. (
SNoL_SNoS ) We take the following as an axiom:
Axiom. (
SNoR_SNoS ) We take the following as an axiom:
Axiom. (
SNoL_I ) We take the following as an axiom:
Axiom. (
SNoR_I ) We take the following as an axiom:
Axiom. (
SNo_eta ) We take the following as an axiom:
Axiom. (
SNoCut_Le ) We take the following as an axiom:
Axiom. (
SNoCut_ext ) We take the following as an axiom:
Axiom. (
ordinal_SNo ) We take the following as an axiom:
Axiom. (
nat_p_SNo ) We take the following as an axiom:
Axiom. (
omega_SNo ) 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. (
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_pos ) We take the following as an axiom:
Axiom. (
eps_SNo_eq ) We take the following as an axiom:
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:
Axiom. (
SNo_ind ) We take the following as an axiom:
Beginning of Section SurrealRecI
Variable F : set → (set → set ) → set
Primitive . The name
SNo_rec_i is a term of type
set → set .
Hypothesis Fr : ∀z, SNo z → ∀g h : set → set , (∀ w ∈ SNoS_ (SNoLev z ) , g w = h w ) → F z g = F z h
End of Section SurrealRecI
Beginning of Section SurrealRecII
Variable F : set → (set → (set → set ) ) → (set → set )
Let G : set → (set → set → (set → set ) ) → set → (set → set ) ≝ λalpha g ⇒ If_iii (ordinal alpha ) (λz : set ⇒ If_ii (z ∈ SNoS_ (ordsucc alpha ) ) (F z (λw ⇒ g (SNoLev w ) w ) ) default ) (λz : set ⇒ default )
Primitive . The name
SNo_rec_ii is a term of type
set → (set → set ) .
Hypothesis Fr : ∀z, SNo z → ∀g h : set → (set → set ) , (∀ w ∈ SNoS_ (SNoLev z ) , g w = h w ) → F z g = F z h
End of Section SurrealRecII
Beginning of Section SurrealRec2
Variable F : set → set → (set → set → set ) → set
Let G : set → (set → set → set ) → set → (set → set ) → set ≝ λw f z g ⇒ F w z (λx y ⇒ if x = w then g y else f x y )
Primitive . The name
SNo_rec2 is a term of type
set → set → set .
Axiom. (
SNo_rec2_eq ) We take the following as an axiom:
End of Section SurrealRec2
Axiom. (
SNoLev_ind ) We take the following as an axiom:
Axiom. (
SNoLev_ind2 ) We take the following as an axiom:
Axiom. (
SNoLev_ind3 ) We take the following as an axiom:
∀P : set → set → set → prop , (∀x y z, SNo x → SNo y → SNo z → (∀ u ∈ SNoS_ (SNoLev x ) , P u y z ) → (∀ v ∈ SNoS_ (SNoLev y ) , P x v z ) → (∀ w ∈ SNoS_ (SNoLev z ) , P x y w ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ v ∈ SNoS_ (SNoLev y ) , P u v z ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ w ∈ SNoS_ (SNoLev z ) , P u y w ) → (∀ v ∈ SNoS_ (SNoLev y ) , ∀ w ∈ SNoS_ (SNoLev z ) , P x v w ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ v ∈ SNoS_ (SNoLev y ) , ∀ w ∈ SNoS_ (SNoLev z ) , P u v w ) → P x y z ) → ∀x y z, SNo x → SNo y → SNo z → P x y z
Axiom. (
SNo_omega ) We take the following as an axiom:
Axiom. (
SNoLt_0_1 ) We take the following as an axiom:
Axiom. (
SNoLt_0_2 ) We take the following as an axiom:
Axiom. (
SNoLt_1_2 ) We take the following as an axiom:
Axiom. (
restr_SNo_ ) We take the following as an axiom:
Axiom. (
restr_SNo ) We take the following as an axiom:
Axiom. (
restr_SNoEq ) We take the following as an axiom:
Beginning of Section SurrealMinus
Primitive . The name
minus_SNo is a term of type
set → set .
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 .
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
set → set → set .
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:
Axiom. (
SNo_add_SNo ) We take the following as an axiom:
Axiom. (
add_SNo_Lt1 ) We take the following as an axiom:
Axiom. (
add_SNo_Le1 ) We take the following as an axiom:
Axiom. (
add_SNo_Lt2 ) We take the following as an axiom:
Axiom. (
add_SNo_Le2 ) We take the following as an axiom:
Axiom. (
add_SNo_Lt3 ) We take the following as an axiom:
Axiom. (
add_SNo_Le3 ) We take the following as an axiom:
Axiom. (
add_SNo_com ) We take the following as an axiom:
Axiom. (
add_SNo_0L ) We take the following as an axiom:
Axiom. (
add_SNo_0R ) We take the following as an axiom:
Axiom. (
minus_SNo_0 ) We take the following as an axiom:
Axiom. (
add_SNo_Lt4 ) We take the following as an axiom:
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 .
Notation . We use
* as an infix operator with priority 355 and which associates to the right corresponding to applying term
mul_SNo .
Proof:
Let x be given.
Let y be given.
Let g and h be given.
We will
prove F x y g = F x y h .
Let w0 be given.
Let w1 be given.
An
exact proof term for the current goal is
SNoL_SNoS x Hx w0 Hw0 .
An
exact proof term for the current goal is
SNoL_SNoS y Hy w1 Hw1 .
We prove the intermediate
claim Lw1b :
SNo w1 .
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim L1aa :
g w0 y = h w0 y .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0 .
An exact proof term for the current goal is Hy .
We prove the intermediate
claim L1ab :
g x w1 = h x w1 .
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lw1 .
We prove the intermediate
claim L1ac :
g w0 w1 = h w0 w1 .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0 .
An exact proof term for the current goal is Lw1b .
rewrite the current goal using L1aa (from left to right).
rewrite the current goal using L1ab (from left to right).
rewrite the current goal using L1ac (from left to right).
Use reflexivity.
Let z0 be given.
Let z1 be given.
An
exact proof term for the current goal is
SNoR_SNoS x Hx z0 Hz0 .
An
exact proof term for the current goal is
SNoR_SNoS y Hy z1 Hz1 .
We prove the intermediate
claim Lz1b :
SNo z1 .
Apply SNoR_E y Hy z1 Hz1 to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim L1ba :
g z0 y = h z0 y .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0 .
An exact proof term for the current goal is Hy .
We prove the intermediate
claim L1bb :
g x z1 = h x z1 .
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lz1 .
We prove the intermediate
claim L1bc :
g z0 z1 = h z0 z1 .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0 .
An exact proof term for the current goal is Lz1b .
rewrite the current goal using L1ba (from left to right).
rewrite the current goal using L1bb (from left to right).
rewrite the current goal using L1bc (from left to right).
Use reflexivity.
Let w0 be given.
Let w1 be given.
An
exact proof term for the current goal is
SNoL_SNoS x Hx w0 Hw0 .
An
exact proof term for the current goal is
SNoR_SNoS y Hy w1 Hw1 .
We prove the intermediate
claim Lw1b :
SNo w1 .
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim L1ca :
g w0 y = h w0 y .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0 .
An exact proof term for the current goal is Hy .
We prove the intermediate
claim L1cb :
g x w1 = h x w1 .
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lw1 .
We prove the intermediate
claim L1cc :
g w0 w1 = h w0 w1 .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lw0 .
An exact proof term for the current goal is Lw1b .
rewrite the current goal using L1ca (from left to right).
rewrite the current goal using L1cb (from left to right).
rewrite the current goal using L1cc (from left to right).
Use reflexivity.
Let z0 be given.
Let z1 be given.
An
exact proof term for the current goal is
SNoR_SNoS x Hx z0 Hz0 .
An
exact proof term for the current goal is
SNoL_SNoS y Hy z1 Hz1 .
We prove the intermediate
claim Lz1b :
SNo z1 .
Apply SNoL_E y Hy z1 Hz1 to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim L1da :
g z0 y = h z0 y .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0 .
An exact proof term for the current goal is Hy .
We prove the intermediate
claim L1db :
g x z1 = h x z1 .
Apply Hgh2 to the current goal.
An exact proof term for the current goal is Lz1 .
We prove the intermediate
claim L1dc :
g z0 z1 = h z0 z1 .
Apply Hgh1 to the current goal.
An exact proof term for the current goal is Lz0 .
An exact proof term for the current goal is Lz1b .
rewrite the current goal using L1da (from left to right).
rewrite the current goal using L1db (from left to right).
rewrite the current goal using L1dc (from left to right).
Use reflexivity.
rewrite the current goal using L1a (from left to right).
rewrite the current goal using L1b (from left to right).
rewrite the current goal using L1c (from left to right).
rewrite the current goal using L1d (from left to right).
Use reflexivity.
An
exact proof term for the current goal is
SNo_rec2_eq F L1 .
∎
Theorem. (
mul_SNo_eq_2 ) The following is provable:
∀x y, SNo x → SNo y → ∀p : prop , (∀L R, (∀u, u ∈ L → (∀q : prop , (∀ w0 ∈ SNoL x , ∀ w1 ∈ SNoL y , u = w0 * y + x * w1 + - w0 * w1 → q ) → (∀ z0 ∈ SNoR x , ∀ z1 ∈ SNoR y , u = z0 * y + x * z1 + - z0 * z1 → q ) → q ) ) → (∀ w0 ∈ SNoL x , ∀ w1 ∈ SNoL y , w0 * y + x * w1 + - w0 * w1 ∈ L ) → (∀ z0 ∈ SNoR x , ∀ z1 ∈ SNoR y , z0 * y + x * z1 + - z0 * z1 ∈ L ) → (∀u, u ∈ R → (∀q : prop , (∀ w0 ∈ SNoL x , ∀ z1 ∈ SNoR y , u = w0 * y + x * z1 + - w0 * z1 → q ) → (∀ z0 ∈ SNoR x , ∀ w1 ∈ SNoL y , u = z0 * y + x * w1 + - z0 * w1 → q ) → q ) ) → (∀ w0 ∈ SNoL x , ∀ z1 ∈ SNoR y , w0 * y + x * z1 + - w0 * z1 ∈ R ) → (∀ z0 ∈ SNoR x , ∀ w1 ∈ SNoL y , z0 * y + x * w1 + - z0 * w1 ∈ R ) → x * y = SNoCut L R → p ) → p
Proof: Let x and y be given.
Assume Hx Hy .
Let p be given.
Assume Hp .
We will prove p .
Apply Hp (Lxy1 ∪ Lxy2 ) (Rxy1 ∪ Rxy2 ) to the current goal.
Let u be given.
Assume Hu .
Let q be given.
Assume Hq1 Hq2 .
Apply binunionE Lxy1 Lxy2 u Hu to the current goal.
Let w0 be given.
Let w1 be given.
We will
prove w0 * y + x * w1 + - w0 * w1 ∈ Lxy1 ∪ Lxy2 .
We will
prove w0 * y + x * w1 + - w0 * w1 ∈ Lxy1 .
We will
prove (w0 ,w1 ) 0 * y + x * w1 + - (w0 ,w1 ) 0 * w1 ∈ Lxy1 .
We will
prove (w0 ,w1 ) 0 * y + x * (w0 ,w1 ) 1 + - (w0 ,w1 ) 0 * (w0 ,w1 ) 1 ∈ Lxy1 .
An exact proof term for the current goal is Hw0 .
An exact proof term for the current goal is Hw1 .
Let z0 be given.
Let z1 be given.
We will
prove z0 * y + x * z1 + - z0 * z1 ∈ Lxy1 ∪ Lxy2 .
We will
prove z0 * y + x * z1 + - z0 * z1 ∈ Lxy2 .
We will
prove (z0 ,z1 ) 0 * y + x * z1 + - (z0 ,z1 ) 0 * z1 ∈ Lxy2 .
We will
prove (z0 ,z1 ) 0 * y + x * (z0 ,z1 ) 1 + - (z0 ,z1 ) 0 * (z0 ,z1 ) 1 ∈ Lxy2 .
An exact proof term for the current goal is Hz0 .
An exact proof term for the current goal is Hz1 .
Let u be given.
Assume Hu .
Let q be given.
Assume Hq1 Hq2 .
Apply binunionE Rxy1 Rxy2 u Hu to the current goal.
Let w0 be given.
Let z1 be given.
We will
prove w0 * y + x * z1 + - w0 * z1 ∈ Rxy1 ∪ Rxy2 .
We will
prove w0 * y + x * z1 + - w0 * z1 ∈ Rxy1 .
We will
prove (w0 ,z1 ) 0 * y + x * z1 + - (w0 ,z1 ) 0 * z1 ∈ Rxy1 .
We will
prove (w0 ,z1 ) 0 * y + x * (w0 ,z1 ) 1 + - (w0 ,z1 ) 0 * (w0 ,z1 ) 1 ∈ Rxy1 .
An exact proof term for the current goal is Hw0 .
An exact proof term for the current goal is Hz1 .
Let z0 be given.
Let w1 be given.
We will
prove z0 * y + x * w1 + - z0 * w1 ∈ Rxy1 ∪ Rxy2 .
We will
prove z0 * y + x * w1 + - z0 * w1 ∈ Rxy2 .
We will
prove (z0 ,w1 ) 0 * y + x * w1 + - (z0 ,w1 ) 0 * w1 ∈ Rxy2 .
We will
prove (z0 ,w1 ) 0 * y + x * (z0 ,w1 ) 1 + - (z0 ,w1 ) 0 * (z0 ,w1 ) 1 ∈ Rxy2 .
An exact proof term for the current goal is Hz0 .
An exact proof term for the current goal is Hw1 .
An
exact proof term for the current goal is
mul_SNo_eq x Hx y Hy .
∎
Theorem. (
mul_SNo_prop_1 ) The following is provable:
∀x, SNo x → ∀y, SNo y → ∀p : prop , (SNo (x * y ) → (∀ u ∈ SNoL x , ∀ v ∈ SNoL y , u * y + x * v < x * y + u * v ) → (∀ u ∈ SNoR x , ∀ v ∈ SNoR y , u * y + x * v < x * y + u * v ) → (∀ u ∈ SNoL x , ∀ v ∈ SNoR y , x * y + u * v < u * y + x * v ) → (∀ u ∈ SNoR x , ∀ v ∈ SNoL y , x * y + u * v < u * y + x * v ) → p ) → p
Proof: Set P to be the term
λx ⇒ ∀y, SNo y → ∀p : prop , (SNo (x * y ) → (∀ u ∈ SNoL x , ∀ v ∈ SNoL y , u * y + x * v < x * y + u * v ) → (∀ u ∈ SNoR x , ∀ v ∈ SNoR y , u * y + x * v < x * y + u * v ) → (∀ u ∈ SNoL x , ∀ v ∈ SNoR y , x * y + u * v < u * y + x * v ) → (∀ u ∈ SNoR x , ∀ v ∈ SNoL y , x * y + u * v < u * y + x * v ) → p ) → p of type
set → prop .
We will
prove ∀x, SNo x → P x .
Let x be given.
We will
prove ∀y, SNo y → Q y .
Let y be given.
Let L and R be given.
Assume HLE HLI1 HLI2 HRE HRI1 HRI2 Hxy .
We prove the intermediate
claim LLx :
∀ u ∈ SNoL x , ∀v, SNo v → SNo (u * v ) .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv .
Apply IHx u (SNoL_SNoS x Hx u Hu ) v Hv to the current goal.
Assume H1 _ _ _ _ .
An exact proof term for the current goal is H1 .
We prove the intermediate
claim LRx :
∀ u ∈ SNoR x , ∀v, SNo v → SNo (u * v ) .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv .
Apply IHx u (SNoR_SNoS x Hx u Hu ) v Hv to the current goal.
Assume H1 _ _ _ _ .
An exact proof term for the current goal is H1 .
We prove the intermediate
claim LLxy :
∀ u ∈ SNoL x , SNo (u * y ) .
Let u be given.
Assume Hu .
An exact proof term for the current goal is LLx u Hu y Hy .
We prove the intermediate
claim LRxy :
∀ u ∈ SNoR x , SNo (u * y ) .
Let u be given.
Assume Hu .
An exact proof term for the current goal is LRx u Hu y Hy .
We prove the intermediate
claim LxLy :
∀ v ∈ SNoL y , SNo (x * v ) .
Let v be given.
Assume Hv .
Apply IHy v (SNoL_SNoS y Hy v Hv ) to the current goal.
Assume H1 _ _ _ _ .
An exact proof term for the current goal is H1 .
We prove the intermediate
claim LxRy :
∀ v ∈ SNoR y , SNo (x * v ) .
Let v be given.
Assume Hv .
Apply IHy v (SNoR_SNoS y Hy v Hv ) to the current goal.
Assume H1 _ _ _ _ .
An exact proof term for the current goal is H1 .
We prove the intermediate
claim LLR1 :
SNoCutP L R .
Apply and3I to the current goal.
Let u be given.
Assume Hu .
Apply HLE u Hu to the current goal.
Let w0 be given.
Assume Hw0 .
Let w1 be given.
Assume Hw1 Huw .
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw1a _ _ .
rewrite the current goal using Huw (from left to right).
We will
prove SNo (w0 * y + x * w1 + - w0 * w1 ) .
We will
prove SNo (w0 * y ) .
An exact proof term for the current goal is LLxy w0 Hw0 .
We will
prove SNo (x * w1 + - w0 * w1 ) .
We will
prove SNo (x * w1 ) .
An exact proof term for the current goal is LxLy w1 Hw1 .
We will
prove SNo (- w0 * w1 ) .
An exact proof term for the current goal is LLx w0 Hw0 w1 Hw1a .
Let z0 be given.
Assume Hz0 .
Let z1 be given.
Assume Hz1 Huz .
Apply SNoR_E y Hy z1 Hz1 to the current goal.
Assume Hz1a _ _ .
rewrite the current goal using Huz (from left to right).
We will
prove SNo (z0 * y + x * z1 + - z0 * z1 ) .
We will
prove SNo (z0 * y ) .
An exact proof term for the current goal is LRxy z0 Hz0 .
We will
prove SNo (x * z1 + - z0 * z1 ) .
We will
prove SNo (x * z1 ) .
An exact proof term for the current goal is LxRy z1 Hz1 .
We will
prove SNo (- z0 * z1 ) .
An exact proof term for the current goal is LRx z0 Hz0 z1 Hz1a .
Let u be given.
Assume Hu .
Apply HRE u Hu to the current goal.
Let w0 be given.
Assume Hw0 .
Let z1 be given.
Assume Hz1 Huw .
Apply SNoR_E y Hy z1 Hz1 to the current goal.
Assume Hz1a _ _ .
rewrite the current goal using Huw (from left to right).
We will
prove SNo (w0 * y + x * z1 + - w0 * z1 ) .
We will
prove SNo (w0 * y ) .
An exact proof term for the current goal is LLxy w0 Hw0 .
We will
prove SNo (x * z1 + - w0 * z1 ) .
We will
prove SNo (x * z1 ) .
An exact proof term for the current goal is LxRy z1 Hz1 .
We will
prove SNo (- w0 * z1 ) .
An exact proof term for the current goal is LLx w0 Hw0 z1 Hz1a .
Let z0 be given.
Assume Hz0 .
Let w1 be given.
Assume Hw1 Huz .
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw1a _ _ .
rewrite the current goal using Huz (from left to right).
We will
prove SNo (z0 * y + x * w1 + - z0 * w1 ) .
We will
prove SNo (z0 * y ) .
An exact proof term for the current goal is LRxy z0 Hz0 .
We will
prove SNo (x * w1 + - z0 * w1 ) .
We will
prove SNo (x * w1 ) .
An exact proof term for the current goal is LxLy w1 Hw1 .
We will
prove SNo (- z0 * w1 ) .
An exact proof term for the current goal is LRx z0 Hz0 w1 Hw1a .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv .
Let u0 be given.
Assume Hu0 .
Let v0 be given.
Assume Hv0 .
Let u1 be given.
Assume Hu1 .
Let v1 be given.
Assume Hv1 .
Assume _ .
Assume _ .
We prove the intermediate
claim Lu0u1 :
SNo (u0 * u1 ) .
Apply IHx u0 Hu0 u1 Hu1c to the current goal.
Assume IHx0 _ _ _ _ .
An exact proof term for the current goal is IHx0 .
We prove the intermediate
claim Lv0v1 :
SNo (v0 * v1 ) .
Apply IHx v0 Hv0 v1 Hv1c to the current goal.
Assume IHx0 _ _ _ _ .
An exact proof term for the current goal is IHx0 .
We prove the intermediate
claim Lu0y :
SNo (u0 * y ) .
Apply IHx u0 Hu0 y Hy to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lxu1 :
SNo (x * u1 ) .
Apply IHy u1 Hu1 to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lv0y :
SNo (v0 * y ) .
Apply IHx v0 Hv0 y Hy to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lxv1 :
SNo (x * v1 ) .
Apply IHy v1 Hv1 to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lu0v1 :
SNo (u0 * v1 ) .
Apply IHx u0 Hu0 v1 Hv1c to the current goal.
Assume IHx0 _ _ _ _ .
An exact proof term for the current goal is IHx0 .
We prove the intermediate
claim Lv0u1 :
SNo (v0 * u1 ) .
Apply IHx v0 Hv0 u1 Hu1c to the current goal.
Assume IHx0 _ _ _ _ .
An exact proof term for the current goal is IHx0 .
Let p be given.
Assume Hp .
Apply Hp Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 to the current goal.
Assume Hue Hve H1 .
rewrite the current goal using Hue (from left to right).
rewrite the current goal using Hve (from left to right).
We prove the intermediate
claim L1 :
(u0 * y + x * u1 + - u0 * u1 ) + (u0 * u1 + v0 * v1 ) = u0 * y + x * u1 + v0 * v1 .
We prove the intermediate
claim L2 :
(v0 * y + x * v1 + - v0 * v1 ) + (u0 * u1 + v0 * v1 ) = v0 * y + x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (u0 * u1 ) (v0 * v1 ) Lu0u1 Lv0v1 (from left to right).
rewrite the current goal using L1 (from left to right).
rewrite the current goal using L2 (from left to right).
An exact proof term for the current goal is H1 .
Apply HLE u Hu to the current goal.
Let u0 be given.
Assume Hu0 .
Let u1 be given.
Assume Hu1 Hue .
Apply SNoL_E x Hx u0 Hu0 to the current goal.
Apply SNoL_E y Hy u1 Hu1 to the current goal.
Apply HRE v Hv to the current goal.
Let v0 be given.
Assume Hv0 .
Let v1 be given.
Assume Hv1 Hve .
Apply SNoL_E x Hx v0 Hv0 to the current goal.
Apply SNoR_E y Hy v1 Hv1 to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv .
Apply Luv Hue Hve to the current goal.
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
We prove the intermediate
claim Lu1v1lt :
u1 < v1 .
An
exact proof term for the current goal is
SNoLt_tra u1 y v1 Hu1a Hy Hv1a Hu1c Hv1c .
We prove the intermediate
claim L3 :
∀ z ∈ SNoL x , x * u1 + z * v1 < z * u1 + x * v1 .
Let z be given.
Assume Hz .
We prove the intermediate
claim Lzu1 :
SNo (z * u1 ) .
An exact proof term for the current goal is LLx z Hz u1 Hu1a .
We prove the intermediate
claim Lzv1 :
SNo (z * v1 ) .
An exact proof term for the current goal is LLx z Hz v1 Hv1a .
Let w be given.
Apply SNoR_E u1 Hu1a w Hwu1 to the current goal.
Apply SNoS_I2 w y Hwu1a Hy to the current goal.
We prove the intermediate
claim Lzw :
SNo (z * w ) .
An exact proof term for the current goal is LLx z Hz w Hwu1a .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
Apply IHy w LwSy to the current goal.
Assume H1 _ _ _ _ .
An exact proof term for the current goal is H1 .
We prove the intermediate
claim L3a :
z * v1 + x * w < x * v1 + z * w .
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ IHy1 _ _ _ .
An exact proof term for the current goal is IHy1 z Hz w Hwv1 .
We prove the intermediate
claim L3b :
x * u1 + z * w < z * u1 + x * w .
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ IHy3 _ .
An exact proof term for the current goal is IHy3 z Hz w Hwu1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop2 (x * u1 ) (z * v1 ) (z * u1 ) (x * v1 ) (z * w ) (x * w ) Lxu1 Lzv1 Lzu1 Lxv1 Lzw Lxw L3b L3a .
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ IHy1 _ _ _ .
rewrite the current goal using
add_SNo_com (x * u1 ) (z * v1 ) Lxu1 Lzv1 (from left to right).
rewrite the current goal using
add_SNo_com (z * u1 ) (x * v1 ) Lzu1 Lxv1 (from left to right).
We will
prove z * v1 + x * u1 < x * v1 + z * u1 .
An exact proof term for the current goal is IHy1 z Hz u1 Hu1v1 .
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ IHy3 _ .
An exact proof term for the current goal is IHy3 z Hz v1 Hv1u1 .
We prove the intermediate
claim L3u0 :
x * u1 + u0 * v1 < u0 * u1 + x * v1 .
An exact proof term for the current goal is L3 u0 Hu0 .
We prove the intermediate
claim L3v0 :
x * u1 + v0 * v1 < v0 * u1 + x * v1 .
An exact proof term for the current goal is L3 v0 Hv0 .
We prove the intermediate
claim L3u0imp :
u0 * y + v0 * v1 < v0 * y + u0 * v1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
Assume H1 .
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0 .
We prove the intermediate
claim L3v0imp :
u0 * y + v0 * u1 < v0 * y + u0 * u1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
Assume H1 .
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
An exact proof term for the current goal is H1 .
We will
prove x * u1 + v0 * v1 < v0 * u1 + x * v1 .
An exact proof term for the current goal is L3v0 .
rewrite the current goal using Hv0u0 (from left to right).
We will
prove u0 * y + (x * u1 + u0 * v1 ) < u0 * y + (x * v1 + u0 * u1 ) .
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0 .
Let z be given.
We prove the intermediate
claim L4 :
u0 * y + z * v1 < z * y + u0 * v1 .
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ IHx3 _ .
An exact proof term for the current goal is IHx3 z Hzu0 v1 Hv1 .
We prove the intermediate
claim L5 :
z * y + v0 * v1 < v0 * y + z * v1 .
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ _ IHx2 _ _ .
An exact proof term for the current goal is IHx2 z Hzv0 v1 Hv1 .
We prove the intermediate
claim Lz :
z ∈ SNoL x .
Apply SNoL_E u0 Hu0a z Hzu0 to the current goal.
Apply SNoL_I x Hx z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra z u0 x Hza Hu0a Hx Hzc Hu0c .
Apply add_SNo_Lt_subprop3c (u0 * y ) (x * u1 ) (v0 * v1 ) (v0 * y ) (x * v1 + u0 * u1 ) (z * v1 ) (z * y ) (u0 * v1 ) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 ) (LLx z Lz v1 Hv1a ) (LLxy z Lz ) Lu0v1 to the current goal.
An exact proof term for the current goal is L4 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0 .
An exact proof term for the current goal is L5 .
Apply L3u0imp to the current goal.
We will
prove u0 * y + v0 * v1 < v0 * y + u0 * v1 .
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ IHx3 _ .
An exact proof term for the current goal is IHx3 v0 Hv0u0 v1 Hv1 .
Apply L3u0imp to the current goal.
We will
prove u0 * y + v0 * v1 < v0 * y + u0 * v1 .
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ _ IHx2 _ _ .
An exact proof term for the current goal is IHx2 u0 Hu0v0 v1 Hv1 .
Let z be given.
We prove the intermediate
claim L6 :
z * y + v0 * u1 < v0 * y + z * u1 .
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ IHx1 _ _ _ .
An exact proof term for the current goal is IHx1 z Hzv0 u1 Hu1 .
We prove the intermediate
claim L7 :
u0 * y + z * u1 < z * y + u0 * u1 .
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ _ IHx4 .
An exact proof term for the current goal is IHx4 z Hzu0 u1 Hu1 .
We prove the intermediate
claim Lz :
z ∈ SNoL x .
Apply SNoL_E v0 Hv0a z Hzv0 to the current goal.
Apply SNoL_I x Hx z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra z v0 x Hza Hv0a Hx Hzc Hv0c .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3d (u0 * y ) (x * u1 + v0 * v1 ) (v0 * y ) (x * v1 ) (u0 * u1 ) (z * u1 ) (z * y ) (v0 * u1 ) Lu0y (SNo_add_SNo (x * u1 ) (v0 * v1 ) Lxu1 Lv0v1 ) Lv0y Lxv1 Lu0u1 (LLx z Lz u1 Hu1a ) (LLxy z Lz ) Lv0u1 L7 L3v0 L6 .
Apply L3v0imp to the current goal.
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ _ IHx4 .
An exact proof term for the current goal is IHx4 v0 Hv0u0 u1 Hu1 .
Apply L3v0imp to the current goal.
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ IHx1 _ _ _ .
An exact proof term for the current goal is IHx1 u0 Hu0v0 u1 Hu1 .
Let v0 be given.
Assume Hv0 .
Let v1 be given.
Assume Hv1 Hve .
Apply SNoR_E x Hx v0 Hv0 to the current goal.
Apply SNoL_E y Hy v1 Hv1 to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv .
Apply Luv Hue Hve to the current goal.
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
We prove the intermediate
claim Lu0v0lt :
u0 < v0 .
An
exact proof term for the current goal is
SNoLt_tra u0 x v0 Hu0a Hx Hv0a Hu0c Hv0c .
We prove the intermediate
claim L3 :
∀ z ∈ SNoL y , u0 * y + v0 * z < v0 * y + u0 * z .
Let z be given.
Assume Hz .
Apply SNoL_E y Hy z Hz to the current goal.
We prove the intermediate
claim Lu0z :
SNo (u0 * z ) .
An exact proof term for the current goal is LLx u0 Hu0 z Hza .
We prove the intermediate
claim Lv0z :
SNo (v0 * z ) .
An exact proof term for the current goal is LRx v0 Hv0 z Hza .
Let w be given.
Apply SNoR_E u0 Hu0a w Hwu0 to the current goal.
An
exact proof term for the current goal is
SNoS_I2 w x Hwu0a Hx LLevwLevx .
We prove the intermediate
claim Lwz :
SNo (w * z ) .
Apply IHx w LwSx z Hza to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lwy :
SNo (w * y ) .
Apply IHx w LwSx y Hy to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim L3a :
u0 * y + w * z < u0 * z + w * y .
rewrite the current goal using
add_SNo_com (u0 * z ) (w * y ) Lu0z Lwy (from left to right).
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ _ IHx4 .
An exact proof term for the current goal is IHx4 w Hwu0 z Hz .
We prove the intermediate
claim L3b :
v0 * z + w * y < v0 * y + w * z .
rewrite the current goal using
add_SNo_com (v0 * z ) (w * y ) Lv0z Lwy (from left to right).
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ IHx1 _ _ _ .
An exact proof term for the current goal is IHx1 w Hwv0 z Hz .
rewrite the current goal using
add_SNo_com (u0 * z ) (v0 * y ) Lu0z Lv0y (from right to left).
An
exact proof term for the current goal is
add_SNo_Lt_subprop2 (u0 * y ) (v0 * z ) (u0 * z ) (v0 * y ) (w * z ) (w * y ) Lu0y Lv0z Lu0z Lv0y Lwz Lwy L3a L3b .
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ IHx1 _ _ _ .
An exact proof term for the current goal is IHx1 u0 Hu0v0 z Hz .
Apply IHx u0 (SNoL_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ _ IHx4 .
An exact proof term for the current goal is IHx4 v0 Hv0u0 z Hz .
We prove the intermediate
claim L3u1 :
u0 * y + v0 * u1 < v0 * y + u0 * u1 .
An exact proof term for the current goal is L3 u1 Hu1 .
We prove the intermediate
claim L3v1 :
u0 * y + v0 * v1 < v0 * y + u0 * v1 .
An exact proof term for the current goal is L3 v1 Hv1 .
We prove the intermediate
claim L3u1imp :
x * u1 + v0 * v1 < v0 * u1 + x * v1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3b (u0 * y ) (x * u1 + v0 * v1 ) (v0 * y ) (x * v1 ) (u0 * u1 ) (v0 * u1 ) Lu0y (SNo_add_SNo (x * u1 ) (v0 * v1 ) Lxu1 Lv0v1 ) Lv0y Lxv1 Lu0u1 Lv0u1 L3u1 .
We prove the intermediate
claim L3v1imp :
x * u1 + u0 * v1 < x * v1 + u0 * u1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3a (u0 * y ) (x * u1 ) (v0 * v1 ) (v0 * y ) (x * v1 + u0 * u1 ) (u0 * v1 ) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 ) Lu0v1 L3v1 .
rewrite the current goal using Hv1u1 (from left to right).
We will
prove u0 * y + x * u1 + v0 * u1 < v0 * y + x * u1 + u0 * u1 .
rewrite the current goal using
add_SNo_com_3_0_1 (u0 * y ) (x * u1 ) (v0 * u1 ) Lu0y Lxu1 Lv0u1 (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (v0 * y ) (x * u1 ) (u0 * u1 ) Lv0y Lxu1 Lu0u1 (from left to right).
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
An exact proof term for the current goal is L3u1 .
Let z be given.
Apply SNoL_E u1 Hu1a z Hzu1 to the current goal.
We prove the intermediate
claim Lz :
z ∈ SNoL y .
Apply SNoL_I y Hy z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra z u1 y Hza Hu1a Hy Hzc Hu1c .
We prove the intermediate
claim L4 :
x * u1 + v0 * z < x * z + v0 * u1 .
rewrite the current goal using
add_SNo_com (x * z ) (v0 * u1 ) (LxLy z Lz ) Lv0u1 (from left to right).
We will
prove x * u1 + v0 * z < v0 * u1 + x * z .
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ _ IHy4 .
An exact proof term for the current goal is IHy4 v0 Hv0 z Hzu1 .
We prove the intermediate
claim L5 :
x * z + v0 * v1 < x * v1 + v0 * z .
rewrite the current goal using
add_SNo_com (x * z ) (v0 * v1 ) (LxLy z Lz ) Lv0v1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ _ IHy2 _ _ .
An exact proof term for the current goal is IHy2 v0 Hv0 z Hzv1 .
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com_3_0_1 (u0 * y ) (x * u1 ) (v0 * v1 ) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (v0 * y ) (x * v1 ) (u0 * u1 ) Lv0y Lxv1 Lu0u1 (from left to right).
We will
prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3c (x * u1 ) (u0 * y ) (v0 * v1 ) (x * v1 ) (v0 * y + u0 * u1 ) (v0 * z ) (x * z ) (v0 * u1 ) Lxu1 Lu0y Lv0v1 Lxv1 (SNo_add_SNo (v0 * y ) (u0 * u1 ) Lv0y Lu0u1 ) (LRx v0 Hv0 z Hza ) (LxLy z Lz ) Lv0u1 L4 L3u1 L5 .
Apply L3u1imp to the current goal.
We will
prove x * u1 + v0 * v1 < v0 * u1 + x * v1 .
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ _ IHy4 .
An exact proof term for the current goal is IHy4 v0 Hv0 v1 Hv1u1 .
Apply L3u1imp to the current goal.
We will
prove x * u1 + v0 * v1 < v0 * u1 + x * v1 .
rewrite the current goal using
add_SNo_com (x * u1 ) (v0 * v1 ) Lxu1 Lv0v1 (from left to right).
rewrite the current goal using
add_SNo_com (v0 * u1 ) (x * v1 ) Lv0u1 Lxv1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ _ IHy2 _ _ .
An exact proof term for the current goal is IHy2 v0 Hv0 u1 Hu1v1 .
Let z be given.
Apply SNoL_E v1 Hv1a z Hzv1 to the current goal.
We prove the intermediate
claim Lz :
z ∈ SNoL y .
Apply SNoL_I y Hy z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra z v1 y Hza Hv1a Hy Hzc Hv1c .
We prove the intermediate
claim L6 :
x * u1 + u0 * z < x * z + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * z ) (u0 * u1 ) (LxLy z Lz ) Lu0u1 (from left to right).
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ IHy3 _ .
An exact proof term for the current goal is IHy3 u0 Hu0 z Hzu1 .
We prove the intermediate
claim L7 :
x * z + u0 * v1 < x * v1 + u0 * z .
rewrite the current goal using
add_SNo_com (x * z ) (u0 * v1 ) (LxLy z Lz ) Lu0v1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ IHy1 _ _ _ .
An exact proof term for the current goal is IHy1 u0 Hu0 z Hzv1 .
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com_3_0_1 (u0 * y ) (x * u1 ) (v0 * v1 ) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (v0 * y ) (x * v1 ) (u0 * u1 ) Lv0y Lxv1 Lu0u1 (from left to right).
We will
prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1 .
Apply add_SNo_Lt_subprop3d (x * u1 ) (u0 * y + v0 * v1 ) (x * v1 ) (v0 * y ) (u0 * u1 ) (u0 * z ) (x * z ) (u0 * v1 ) Lxu1 (SNo_add_SNo (u0 * y ) (v0 * v1 ) Lu0y Lv0v1 ) Lxv1 Lv0y Lu0u1 (LLx u0 Hu0 z Hza ) (LxLy z Lz ) Lu0v1 L6 to the current goal.
We will
prove u0 * y + v0 * v1 < u0 * v1 + v0 * y .
rewrite the current goal using
add_SNo_com (u0 * v1 ) (v0 * y ) Lu0v1 Lv0y (from left to right).
An exact proof term for the current goal is L3v1 .
An exact proof term for the current goal is L7 .
Apply L3v1imp to the current goal.
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
Apply IHy u1 (SNoL_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ IHy3 _ .
An exact proof term for the current goal is IHy3 u0 Hu0 v1 Hv1u1 .
Apply L3v1imp to the current goal.
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * u1 ) (u0 * v1 ) Lxu1 Lu0v1 (from left to right).
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ IHy1 _ _ _ .
An exact proof term for the current goal is IHy1 u0 Hu0 u1 Hu1v1 .
Let u0 be given.
Assume Hu0 .
Let u1 be given.
Assume Hu1 Hue .
Apply SNoR_E x Hx u0 Hu0 to the current goal.
Apply SNoR_E y Hy u1 Hu1 to the current goal.
Apply HRE v Hv to the current goal.
Let v0 be given.
Assume Hv0 .
Let v1 be given.
Assume Hv1 Hve .
Apply SNoL_E x Hx v0 Hv0 to the current goal.
Apply SNoR_E y Hy v1 Hv1 to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv .
Apply Luv Hue Hve to the current goal.
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
We prove the intermediate
claim Lv0u0lt :
v0 < u0 .
An
exact proof term for the current goal is
SNoLt_tra v0 x u0 Hv0a Hx Hu0a Hv0c Hu0c .
We prove the intermediate
claim L3 :
∀ z ∈ SNoR y , u0 * y + v0 * z < v0 * y + u0 * z .
Let z be given.
Assume Hz .
Apply SNoR_E y Hy z Hz to the current goal.
We prove the intermediate
claim Lu0z :
SNo (u0 * z ) .
An exact proof term for the current goal is LRx u0 Hu0 z Hza .
We prove the intermediate
claim Lv0z :
SNo (v0 * z ) .
An exact proof term for the current goal is LLx v0 Hv0 z Hza .
Let w be given.
Apply SNoL_E u0 Hu0a w Hwu0 to the current goal.
An
exact proof term for the current goal is
SNoS_I2 w x Hwu0a Hx LLevwLevx .
We prove the intermediate
claim Lwz :
SNo (w * z ) .
Apply IHx w LwSx z Hza to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lwy :
SNo (w * y ) .
Apply IHx w LwSx y Hy to the current goal.
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim L3a :
u0 * y + w * z < u0 * z + w * y .
rewrite the current goal using
add_SNo_com (u0 * z ) (w * y ) Lu0z Lwy (from left to right).
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ IHx3 _ .
An exact proof term for the current goal is IHx3 w Hwu0 z Hz .
We prove the intermediate
claim L3b :
v0 * z + w * y < v0 * y + w * z .
rewrite the current goal using
add_SNo_com (v0 * z ) (w * y ) Lv0z Lwy (from left to right).
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ _ IHx2 _ _ .
An exact proof term for the current goal is IHx2 w Hwv0 z Hz .
rewrite the current goal using
add_SNo_com (v0 * y ) (u0 * z ) Lv0y Lu0z (from left to right).
An
exact proof term for the current goal is
add_SNo_Lt_subprop2 (u0 * y ) (v0 * z ) (u0 * z ) (v0 * y ) (w * z ) (w * y ) Lu0y Lv0z Lu0z Lv0y Lwz Lwy L3a L3b .
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ IHx3 _ .
An exact proof term for the current goal is IHx3 v0 Hv0u0 z Hz .
Apply IHx v0 (SNoL_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ _ IHx2 _ _ .
An exact proof term for the current goal is IHx2 u0 Hu0v0 z Hz .
We prove the intermediate
claim L3u1 :
u0 * y + v0 * u1 < v0 * y + u0 * u1 .
An exact proof term for the current goal is L3 u1 Hu1 .
We prove the intermediate
claim L3v1 :
u0 * y + v0 * v1 < v0 * y + u0 * v1 .
An exact proof term for the current goal is L3 v1 Hv1 .
We prove the intermediate
claim L3u1imp :
x * u1 + v0 * v1 < v0 * u1 + x * v1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3b (u0 * y ) (x * u1 + v0 * v1 ) (v0 * y ) (x * v1 ) (u0 * u1 ) (v0 * u1 ) Lu0y (SNo_add_SNo (x * u1 ) (v0 * v1 ) Lxu1 Lv0v1 ) Lv0y Lxv1 Lu0u1 Lv0u1 L3u1 .
We prove the intermediate
claim L3v1imp :
x * u1 + u0 * v1 < x * v1 + u0 * u1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3a (u0 * y ) (x * u1 ) (v0 * v1 ) (v0 * y ) (x * v1 + u0 * u1 ) (u0 * v1 ) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 ) Lu0v1 L3v1 .
rewrite the current goal using Hv1u1 (from left to right).
We will
prove u0 * y + x * u1 + v0 * u1 < v0 * y + x * u1 + u0 * u1 .
rewrite the current goal using
add_SNo_com_3_0_1 (u0 * y ) (x * u1 ) (v0 * u1 ) Lu0y Lxu1 Lv0u1 (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (v0 * y ) (x * u1 ) (u0 * u1 ) Lv0y Lxu1 Lu0u1 (from left to right).
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
An exact proof term for the current goal is L3u1 .
Let z be given.
Apply SNoR_E v1 Hv1a z Hzv1 to the current goal.
We prove the intermediate
claim Lz :
z ∈ SNoR y .
Apply SNoR_I y Hy z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra y v1 z Hy Hv1a Hza Hv1c Hzc .
We prove the intermediate
claim L4 :
x * u1 + u0 * z < x * z + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * z ) (u0 * u1 ) (LxRy z Lz ) Lu0u1 (from left to right).
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ _ IHy4 .
An exact proof term for the current goal is IHy4 u0 Hu0 z Hzu1 .
We prove the intermediate
claim L5 :
x * z + u0 * v1 < x * v1 + u0 * z .
rewrite the current goal using
add_SNo_com (x * z ) (u0 * v1 ) (LxRy z Lz ) Lu0v1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ _ IHy2 _ _ .
An exact proof term for the current goal is IHy2 u0 Hu0 z Hzv1 .
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com_3_0_1 (u0 * y ) (x * u1 ) (v0 * v1 ) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (v0 * y ) (x * v1 ) (u0 * u1 ) Lv0y Lxv1 Lu0u1 (from left to right).
We will
prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1 .
Apply add_SNo_Lt_subprop3d (x * u1 ) (u0 * y + v0 * v1 ) (x * v1 ) (v0 * y ) (u0 * u1 ) (u0 * z ) (x * z ) (u0 * v1 ) Lxu1 (SNo_add_SNo (u0 * y ) (v0 * v1 ) Lu0y Lv0v1 ) Lxv1 Lv0y Lu0u1 (LRx u0 Hu0 z Hza ) (LxRy z Lz ) Lu0v1 L4 to the current goal.
We will
prove u0 * y + v0 * v1 < u0 * v1 + v0 * y .
rewrite the current goal using
add_SNo_com (u0 * v1 ) (v0 * y ) Lu0v1 Lv0y (from left to right).
An exact proof term for the current goal is L3v1 .
An exact proof term for the current goal is L5 .
Apply L3v1imp to the current goal.
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
We will
prove x * u1 + u0 * v1 < u0 * u1 + x * v1 .
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ _ IHy4 .
An exact proof term for the current goal is IHy4 u0 Hu0 v1 Hv1u1 .
Apply L3v1imp to the current goal.
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * u1 ) (u0 * v1 ) Lxu1 Lu0v1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ _ IHy2 _ _ .
An exact proof term for the current goal is IHy2 u0 Hu0 u1 Hu1v1 .
Let z be given.
Apply SNoR_E u1 Hu1a z Hzu1 to the current goal.
We prove the intermediate
claim Lz :
z ∈ SNoR y .
Apply SNoR_I y Hy z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra y u1 z Hy Hu1a Hza Hu1c Hzc .
We prove the intermediate
claim L6 :
x * u1 + v0 * z < x * z + v0 * u1 .
rewrite the current goal using
add_SNo_com (x * z ) (v0 * u1 ) (LxRy z Lz ) Lv0u1 (from left to right).
We will
prove x * u1 + v0 * z < v0 * u1 + x * z .
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ IHy3 _ .
An exact proof term for the current goal is IHy3 v0 Hv0 z Hzu1 .
We prove the intermediate
claim L7 :
x * z + v0 * v1 < x * v1 + v0 * z .
rewrite the current goal using
add_SNo_com (x * z ) (v0 * v1 ) (LxRy z Lz ) Lv0v1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ IHy1 _ _ _ .
An exact proof term for the current goal is IHy1 v0 Hv0 z Hzv1 .
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com_3_0_1 (u0 * y ) (x * u1 ) (v0 * v1 ) Lu0y Lxu1 Lv0v1 (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (v0 * y ) (x * v1 ) (u0 * u1 ) Lv0y Lxv1 Lu0u1 (from left to right).
We will
prove x * u1 + u0 * y + v0 * v1 < x * v1 + v0 * y + u0 * u1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3c (x * u1 ) (u0 * y ) (v0 * v1 ) (x * v1 ) (v0 * y + u0 * u1 ) (v0 * z ) (x * z ) (v0 * u1 ) Lxu1 Lu0y Lv0v1 Lxv1 (SNo_add_SNo (v0 * y ) (u0 * u1 ) Lv0y Lu0u1 ) (LLx v0 Hv0 z Hza ) (LxRy z Lz ) Lv0u1 L6 L3u1 L7 .
Apply L3u1imp to the current goal.
We will
prove x * u1 + v0 * v1 < v0 * u1 + x * v1 .
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ IHy3 _ .
An exact proof term for the current goal is IHy3 v0 Hv0 v1 Hv1u1 .
Apply L3u1imp to the current goal.
We will
prove x * u1 + v0 * v1 < v0 * u1 + x * v1 .
rewrite the current goal using
add_SNo_com (x * u1 ) (v0 * v1 ) Lxu1 Lv0v1 (from left to right).
rewrite the current goal using
add_SNo_com (v0 * u1 ) (x * v1 ) Lv0u1 Lxv1 (from left to right).
Apply IHy v1 (SNoR_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ IHy1 _ _ _ .
An exact proof term for the current goal is IHy1 v0 Hv0 u1 Hu1v1 .
Let v0 be given.
Assume Hv0 .
Let v1 be given.
Assume Hv1 Hve .
Apply SNoR_E x Hx v0 Hv0 to the current goal.
Apply SNoL_E y Hy v1 Hv1 to the current goal.
Assume Lu0y Lxu1 Lu0u1 Lv0y Lxv1 Lv0v1 Lu0v1 Lv0u1 Luv .
Apply Luv Hue Hve to the current goal.
We will
prove u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
We prove the intermediate
claim Lv1u1lt :
v1 < u1 .
An
exact proof term for the current goal is
SNoLt_tra v1 y u1 Hv1a Hy Hu1a Hv1c Hu1c .
We prove the intermediate
claim L3 :
∀ z ∈ SNoR x , x * u1 + z * v1 < z * u1 + x * v1 .
Let z be given.
Assume Hz .
We prove the intermediate
claim Lzu1 :
SNo (z * u1 ) .
An exact proof term for the current goal is LRx z Hz u1 Hu1a .
We prove the intermediate
claim Lzv1 :
SNo (z * v1 ) .
An exact proof term for the current goal is LRx z Hz v1 Hv1a .
Let w be given.
Apply SNoL_E u1 Hu1a w Hwu1 to the current goal.
Apply SNoS_I2 w y Hwu1a Hy to the current goal.
We prove the intermediate
claim Lzw :
SNo (z * w ) .
An exact proof term for the current goal is LRx z Hz w Hwu1a .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
Apply IHy w LwSy to the current goal.
Assume H1 _ _ _ _ .
An exact proof term for the current goal is H1 .
We prove the intermediate
claim L3a :
z * v1 + x * w < x * v1 + z * w .
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ _ IHy2 _ _ .
An exact proof term for the current goal is IHy2 z Hz w Hwv1 .
We prove the intermediate
claim L3b :
x * u1 + z * w < z * u1 + x * w .
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ _ IHy4 .
An exact proof term for the current goal is IHy4 z Hz w Hwu1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop2 (x * u1 ) (z * v1 ) (z * u1 ) (x * v1 ) (z * w ) (x * w ) Lxu1 Lzv1 Lzu1 Lxv1 Lzw Lxw L3b L3a .
Apply IHy u1 (SNoR_SNoS y Hy u1 Hu1 ) to the current goal.
Assume _ _ _ _ IHy4 .
An exact proof term for the current goal is IHy4 z Hz v1 Hv1u1 .
Apply IHy v1 (SNoL_SNoS y Hy v1 Hv1 ) to the current goal.
Assume _ _ IHy2 _ _ .
rewrite the current goal using
add_SNo_com (x * u1 ) (z * v1 ) Lxu1 Lzv1 (from left to right).
rewrite the current goal using
add_SNo_com (z * u1 ) (x * v1 ) Lzu1 Lxv1 (from left to right).
We will
prove z * v1 + x * u1 < x * v1 + z * u1 .
An exact proof term for the current goal is IHy2 z Hz u1 Hu1v1 .
We prove the intermediate
claim L3u0 :
x * u1 + u0 * v1 < u0 * u1 + x * v1 .
An exact proof term for the current goal is L3 u0 Hu0 .
We prove the intermediate
claim L3v0 :
x * u1 + v0 * v1 < v0 * u1 + x * v1 .
An exact proof term for the current goal is L3 v0 Hv0 .
We prove the intermediate
claim L3u0imp :
u0 * y + v0 * v1 < v0 * y + u0 * v1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
Assume H1 .
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0 .
We prove the intermediate
claim L3v0imp :
u0 * y + v0 * u1 < v0 * y + u0 * u1 → u0 * y + x * u1 + v0 * v1 < v0 * y + x * v1 + u0 * u1 .
Assume H1 .
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
An exact proof term for the current goal is H1 .
We will
prove x * u1 + v0 * v1 < v0 * u1 + x * v1 .
An exact proof term for the current goal is L3v0 .
rewrite the current goal using Hv0u0 (from left to right).
We will
prove u0 * y + (x * u1 + u0 * v1 ) < u0 * y + (x * v1 + u0 * u1 ) .
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0 .
Let z be given.
Apply SNoR_E v0 Hv0a z Hzv0 to the current goal.
We prove the intermediate
claim Lz :
z ∈ SNoR x .
Apply SNoR_I x Hx z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra x v0 z Hx Hv0a Hza Hv0c Hzc .
We prove the intermediate
claim L4 :
u0 * y + z * u1 < z * y + u0 * u1 .
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ IHx3 _ .
An exact proof term for the current goal is IHx3 z Hzu0 u1 Hu1 .
We prove the intermediate
claim L5 :
z * y + v0 * u1 < v0 * y + z * u1 .
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ _ IHx2 _ _ .
An exact proof term for the current goal is IHx2 z Hzv0 u1 Hu1 .
An
exact proof term for the current goal is
add_SNo_Lt_subprop3d (u0 * y ) (x * u1 + v0 * v1 ) (v0 * y ) (x * v1 ) (u0 * u1 ) (z * u1 ) (z * y ) (v0 * u1 ) Lu0y (SNo_add_SNo (x * u1 ) (v0 * v1 ) Lxu1 Lv0v1 ) Lv0y Lxv1 Lu0u1 (LRx z Lz u1 Hu1a ) (LRxy z Lz ) Lv0u1 L4 L3v0 L5 .
Apply L3v0imp to the current goal.
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ IHx3 _ .
An exact proof term for the current goal is IHx3 v0 Hv0u0 u1 Hu1 .
Apply L3v0imp to the current goal.
We will
prove u0 * y + v0 * u1 < v0 * y + u0 * u1 .
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ _ IHx2 _ _ .
An exact proof term for the current goal is IHx2 u0 Hu0v0 u1 Hu1 .
Let z be given.
Apply SNoR_E u0 Hu0a z Hzu0 to the current goal.
We prove the intermediate
claim Lz :
z ∈ SNoR x .
Apply SNoR_I x Hx z Hza to the current goal.
An
exact proof term for the current goal is
SNoLt_tra x u0 z Hx Hu0a Hza Hu0c Hzc .
We prove the intermediate
claim L6 :
u0 * y + z * v1 < z * y + u0 * v1 .
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ _ IHx4 .
An exact proof term for the current goal is IHx4 z Hzu0 v1 Hv1 .
We prove the intermediate
claim L7 :
z * y + v0 * v1 < v0 * y + z * v1 .
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ IHx1 _ _ _ .
An exact proof term for the current goal is IHx1 z Hzv0 v1 Hv1 .
Apply add_SNo_Lt_subprop3c (u0 * y ) (x * u1 ) (v0 * v1 ) (v0 * y ) (x * v1 + u0 * u1 ) (z * v1 ) (z * y ) (u0 * v1 ) Lu0y Lxu1 Lv0v1 Lv0y (SNo_add_SNo (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 ) (LRx z Lz v1 Hv1a ) (LRxy z Lz ) Lu0v1 L6 to the current goal.
We will
prove x * u1 + u0 * v1 < x * v1 + u0 * u1 .
rewrite the current goal using
add_SNo_com (x * v1 ) (u0 * u1 ) Lxv1 Lu0u1 (from left to right).
An exact proof term for the current goal is L3u0 .
We will
prove z * y + v0 * v1 < v0 * y + z * v1 .
An exact proof term for the current goal is L7 .
Apply L3u0imp to the current goal.
We will
prove u0 * y + v0 * v1 < v0 * y + u0 * v1 .
Apply IHx u0 (SNoR_SNoS x Hx u0 Hu0 ) y Hy to the current goal.
Assume _ _ _ _ IHx4 .
An exact proof term for the current goal is IHx4 v0 Hv0u0 v1 Hv1 .
Apply L3u0imp to the current goal.
We will
prove u0 * y + v0 * v1 < v0 * y + u0 * v1 .
Apply IHx v0 (SNoR_SNoS x Hx v0 Hv0 ) y Hy to the current goal.
Assume _ IHx1 _ _ _ .
An exact proof term for the current goal is IHx1 u0 Hu0v0 v1 Hv1 .
We prove the intermediate
claim Lxy :
SNo (x * y ) .
rewrite the current goal using Hxy (from left to right).
Let p be given.
Assume Hp .
We will prove p .
Apply Hp to the current goal.
An exact proof term for the current goal is Lxy .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We will
prove u * y + x * v < x * y + u * v .
We prove the intermediate
claim L1 :
u * y + x * v + - u * v < x * y .
rewrite the current goal using Hxy (from left to right).
We will
prove u * y + x * v + - u * v ∈ L .
An exact proof term for the current goal is HLI1 u Hu v Hv .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (x * v ) (LLx u Hu y Hy ) (LxLy v Hv ) .
An exact proof term for the current goal is LLx u Hu v Hva .
An exact proof term for the current goal is Lxy .
We will
prove (u * y + x * v ) + - u * v < x * y .
rewrite the current goal using
add_SNo_assoc (u * y ) (x * v ) (- (u * v ) ) (LLx u Hu y Hy ) (LxLy v Hv ) (SNo_minus_SNo (u * v ) (LLx u Hu v Hva ) ) (from right to left).
An exact proof term for the current goal is L1 .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We will
prove u * y + x * v < x * y + u * v .
We prove the intermediate
claim L1 :
u * y + x * v + - u * v < x * y .
rewrite the current goal using Hxy (from left to right).
We will
prove u * y + x * v + - u * v ∈ L .
An exact proof term for the current goal is HLI2 u Hu v Hv .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (x * v ) (LRx u Hu y Hy ) (LxRy v Hv ) .
An exact proof term for the current goal is LRx u Hu v Hva .
An exact proof term for the current goal is Lxy .
We will
prove (u * y + x * v ) + - u * v < x * y .
rewrite the current goal using
add_SNo_assoc (u * y ) (x * v ) (- (u * v ) ) (LRx u Hu y Hy ) (LxRy v Hv ) (SNo_minus_SNo (u * v ) (LRx u Hu v Hva ) ) (from right to left).
An exact proof term for the current goal is L1 .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We will
prove x * y + u * v < u * y + x * v .
We prove the intermediate
claim L1 :
x * y < u * y + x * v + - u * v .
rewrite the current goal using Hxy (from left to right).
We will
prove u * y + x * v + - u * v ∈ R .
An exact proof term for the current goal is HRI1 u Hu v Hv .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (x * v ) (LLx u Hu y Hy ) (LxRy v Hv ) .
An exact proof term for the current goal is LLx u Hu v Hva .
An exact proof term for the current goal is Lxy .
We will
prove x * y < (u * y + x * v ) + - u * v .
rewrite the current goal using
add_SNo_assoc (u * y ) (x * v ) (- (u * v ) ) (LLx u Hu y Hy ) (LxRy v Hv ) (SNo_minus_SNo (u * v ) (LLx u Hu v Hva ) ) (from right to left).
An exact proof term for the current goal is L1 .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We will
prove x * y + u * v < u * y + x * v .
We prove the intermediate
claim L1 :
x * y < u * y + x * v + - u * v .
rewrite the current goal using Hxy (from left to right).
We will
prove u * y + x * v + - u * v ∈ R .
An exact proof term for the current goal is HRI2 u Hu v Hv .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (x * v ) (LRx u Hu y Hy ) (LxLy v Hv ) .
An exact proof term for the current goal is LRx u Hu v Hva .
An exact proof term for the current goal is Lxy .
We will
prove x * y < (u * y + x * v ) + - u * v .
rewrite the current goal using
add_SNo_assoc (u * y ) (x * v ) (- (u * v ) ) (LRx u Hu y Hy ) (LxLy v Hv ) (SNo_minus_SNo (u * v ) (LRx u Hu v Hva ) ) (from right to left).
An exact proof term for the current goal is L1 .
∎
Proof: Let x and y be given.
Assume Hx Hy .
Assume H _ _ _ _ .
An exact proof term for the current goal is H .
∎
Proof: Let x, y, u and v be given.
Assume Hx Hy Hu Hv .
An exact proof term for the current goal is Hu .
An exact proof term for the current goal is Hy .
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hv .
An exact proof term for the current goal is Hu .
An exact proof term for the current goal is Hv .
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz .
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hy .
An exact proof term for the current goal is Hz .
∎
Theorem. (
mul_SNo_eq_3 ) The following is provable:
∀x y, SNo x → SNo y → ∀p : prop , (∀L R, SNoCutP L R → (∀u, u ∈ L → (∀q : prop , (∀ w0 ∈ SNoL x , ∀ w1 ∈ SNoL y , u = w0 * y + x * w1 + - w0 * w1 → q ) → (∀ z0 ∈ SNoR x , ∀ z1 ∈ SNoR y , u = z0 * y + x * z1 + - z0 * z1 → q ) → q ) ) → (∀ w0 ∈ SNoL x , ∀ w1 ∈ SNoL y , w0 * y + x * w1 + - w0 * w1 ∈ L ) → (∀ z0 ∈ SNoR x , ∀ z1 ∈ SNoR y , z0 * y + x * z1 + - z0 * z1 ∈ L ) → (∀u, u ∈ R → (∀q : prop , (∀ w0 ∈ SNoL x , ∀ z1 ∈ SNoR y , u = w0 * y + x * z1 + - w0 * z1 → q ) → (∀ z0 ∈ SNoR x , ∀ w1 ∈ SNoL y , u = z0 * y + x * w1 + - z0 * w1 → q ) → q ) ) → (∀ w0 ∈ SNoL x , ∀ z1 ∈ SNoR y , w0 * y + x * z1 + - w0 * z1 ∈ R ) → (∀ z0 ∈ SNoR x , ∀ w1 ∈ SNoL y , z0 * y + x * w1 + - z0 * w1 ∈ R ) → x * y = SNoCut L R → p ) → p
Proof: Let x and y be given.
Assume Hx Hy .
Let p be given.
Assume Hp .
Let L and R be given.
Assume HLE HLI1 HLI2 HRE HRI1 HRI2 He .
We prove the intermediate
claim LLR :
SNoCutP L R .
Apply and3I to the current goal.
Let w be given.
Assume Hw .
Apply HLE w Hw to the current goal.
Let u be given.
Assume Hu .
Let v be given.
Assume Hv Hwe .
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using Hwe (from left to right).
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hy .
An exact proof term for the current goal is Hua .
An exact proof term for the current goal is Hva .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv Hwe .
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using Hwe (from left to right).
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hy .
An exact proof term for the current goal is Hua .
An exact proof term for the current goal is Hva .
Let z be given.
Assume Hz .
Apply HRE z Hz to the current goal.
Let u be given.
Assume Hu .
Let v be given.
Assume Hv Hze .
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using Hze (from left to right).
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hy .
An exact proof term for the current goal is Hua .
An exact proof term for the current goal is Hva .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv Hze .
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using Hze (from left to right).
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hy .
An exact proof term for the current goal is Hua .
An exact proof term for the current goal is Hva .
Let w be given.
Assume Hw .
Let z be given.
Assume Hz .
Apply HLE w Hw to the current goal.
Let u be given.
Assume Hu .
Let v be given.
Assume Hv Hwe .
Apply SNoL_E x Hx u Hu to the current goal.
Assume _ _ .
Apply SNoL_E y Hy v Hv to the current goal.
Assume _ _ .
We prove the intermediate
claim Luy :
SNo (u * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo u y Hua Hy .
We prove the intermediate
claim Lxv :
SNo (x * v ) .
An
exact proof term for the current goal is
SNo_mul_SNo x v Hx Hva .
We prove the intermediate
claim Luv :
SNo (u * v ) .
An
exact proof term for the current goal is
SNo_mul_SNo u v Hua Hva .
Apply HRE z Hz to the current goal.
Let u' be given.
Assume Hu' .
Let v' be given.
Assume Hv' Hze .
Apply SNoL_E x Hx u' Hu' to the current goal.
Assume _ _ .
Apply SNoR_E y Hy v' Hv' to the current goal.
Assume _ _ .
We prove the intermediate
claim Lu'y :
SNo (u' * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' y Hu'a Hy .
We prove the intermediate
claim Lxv' :
SNo (x * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo x v' Hx Hv'a .
We prove the intermediate
claim Lu'v' :
SNo (u' * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' v' Hu'a Hv'a .
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will
prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v' .
We will
prove u * y + x * v + u' * v' < u' * y + x * v' + u * v .
We will
prove u * y + x * v + u' * v' < x * y + u * v + u' * v' .
We will
prove u * y + x * v < x * y + u * v .
An exact proof term for the current goal is Hxy1 u Hu v Hv .
We will
prove x * y + u * v + u' * v' < u' * y + x * v' + u * v .
Apply add_SNo_3_2_3_Lt1 (x * y ) (u' * v' ) (u' * y ) (x * v' ) (u * v ) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will
prove u' * v' + x * y < u' * y + x * v' .
rewrite the current goal using
add_SNo_com (u' * v' ) (x * y ) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy3 u' Hu' v' Hv' .
Let u' be given.
Assume Hu' .
Let v' be given.
Assume Hv' Hze .
Apply SNoR_E x Hx u' Hu' to the current goal.
Assume _ _ .
Apply SNoL_E y Hy v' Hv' to the current goal.
Assume _ _ .
We prove the intermediate
claim Lu'y :
SNo (u' * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' y Hu'a Hy .
We prove the intermediate
claim Lxv' :
SNo (x * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo x v' Hx Hv'a .
We prove the intermediate
claim Lu'v' :
SNo (u' * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' v' Hu'a Hv'a .
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will
prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v' .
We will
prove u * y + x * v + u' * v' < u' * y + x * v' + u * v .
We will
prove u * y + x * v + u' * v' < x * y + u * v + u' * v' .
We will
prove u * y + x * v < x * y + u * v .
An exact proof term for the current goal is Hxy1 u Hu v Hv .
We will
prove x * y + u * v + u' * v' < u' * y + x * v' + u * v .
Apply add_SNo_3_2_3_Lt1 (x * y ) (u' * v' ) (u' * y ) (x * v' ) (u * v ) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will
prove u' * v' + x * y < u' * y + x * v' .
rewrite the current goal using
add_SNo_com (u' * v' ) (x * y ) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy4 u' Hu' v' Hv' .
Let u be given.
Assume Hu .
Let v be given.
Assume Hv Hwe .
Apply SNoR_E x Hx u Hu to the current goal.
Assume _ _ .
Apply SNoR_E y Hy v Hv to the current goal.
Assume _ _ .
We prove the intermediate
claim Luy :
SNo (u * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo u y Hua Hy .
We prove the intermediate
claim Lxv :
SNo (x * v ) .
An
exact proof term for the current goal is
SNo_mul_SNo x v Hx Hva .
We prove the intermediate
claim Luv :
SNo (u * v ) .
An
exact proof term for the current goal is
SNo_mul_SNo u v Hua Hva .
Apply HRE z Hz to the current goal.
Let u' be given.
Assume Hu' .
Let v' be given.
Assume Hv' Hze .
Apply SNoL_E x Hx u' Hu' to the current goal.
Assume _ _ .
Apply SNoR_E y Hy v' Hv' to the current goal.
Assume _ _ .
We prove the intermediate
claim Lu'y :
SNo (u' * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' y Hu'a Hy .
We prove the intermediate
claim Lxv' :
SNo (x * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo x v' Hx Hv'a .
We prove the intermediate
claim Lu'v' :
SNo (u' * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' v' Hu'a Hv'a .
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will
prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v' .
We will
prove u * y + x * v + u' * v' < u' * y + x * v' + u * v .
We will
prove u * y + x * v + u' * v' < x * y + u * v + u' * v' .
We will
prove u * y + x * v < x * y + u * v .
An exact proof term for the current goal is Hxy2 u Hu v Hv .
We will
prove x * y + u * v + u' * v' < u' * y + x * v' + u * v .
Apply add_SNo_3_2_3_Lt1 (x * y ) (u' * v' ) (u' * y ) (x * v' ) (u * v ) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will
prove u' * v' + x * y < u' * y + x * v' .
rewrite the current goal using
add_SNo_com (u' * v' ) (x * y ) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy3 u' Hu' v' Hv' .
Let u' be given.
Assume Hu' .
Let v' be given.
Assume Hv' Hze .
Apply SNoR_E x Hx u' Hu' to the current goal.
Assume _ _ .
Apply SNoL_E y Hy v' Hv' to the current goal.
Assume _ _ .
We prove the intermediate
claim Lu'y :
SNo (u' * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' y Hu'a Hy .
We prove the intermediate
claim Lxv' :
SNo (x * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo x v' Hx Hv'a .
We prove the intermediate
claim Lu'v' :
SNo (u' * v' ) .
An
exact proof term for the current goal is
SNo_mul_SNo u' v' Hu'a Hv'a .
rewrite the current goal using Hwe (from left to right).
rewrite the current goal using Hze (from left to right).
We will
prove u * y + x * v + - u * v < u' * y + x * v' + - u' * v' .
We will
prove u * y + x * v + u' * v' < u' * y + x * v' + u * v .
We will
prove u * y + x * v + u' * v' < x * y + u * v + u' * v' .
We will
prove u * y + x * v < x * y + u * v .
An exact proof term for the current goal is Hxy2 u Hu v Hv .
We will
prove x * y + u * v + u' * v' < u' * y + x * v' + u * v .
Apply add_SNo_3_2_3_Lt1 (x * y ) (u' * v' ) (u' * y ) (x * v' ) (u * v ) Hxy Lu'v' Lu'y Lxv' Luv to the current goal.
We will
prove u' * v' + x * y < u' * y + x * v' .
rewrite the current goal using
add_SNo_com (u' * v' ) (x * y ) Lu'v' Hxy (from left to right).
An exact proof term for the current goal is Hxy4 u' Hu' v' Hv' .
An exact proof term for the current goal is Hp L R LLR HLE HLI1 HLI2 HRE HRI1 HRI2 He .
∎
Proof: Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy .
We prove the intermediate
claim Luyxv :
SNo (u * y + x * v ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (x * v ) Huy Hxv .
We prove the intermediate
claim Lxyuv :
SNo (x * y + u * v ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (u * v ) Hxy Huv .
Apply SNoLtE u x Hu Hx Hux to the current goal.
Let w be given.
Assume Hw2a Hw2b .
We prove the intermediate
claim Lwx :
w ∈ SNoL x .
An
exact proof term for the current goal is
SNoL_I x Hx w Hw1 Hw2b Hw6 .
We prove the intermediate
claim Lwu :
w ∈ SNoR u .
An
exact proof term for the current goal is
SNoR_I u Hu w Hw1 Hw2a Hw5 .
We prove the intermediate
claim Lwy :
SNo (w * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo w y Hw1 Hy .
We prove the intermediate
claim Lwv :
SNo (w * v ) .
An
exact proof term for the current goal is
SNo_mul_SNo w v Hw1 Hv .
We prove the intermediate
claim Lwyxv :
SNo (w * y + x * v ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * y ) (x * v ) Lwy Hxv .
We prove the intermediate
claim Lwyuv :
SNo (w * y + u * v ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * y ) (u * v ) Lwy Huv .
We prove the intermediate
claim Lxywv :
SNo (x * y + w * v ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (w * v ) Hxy Lwv .
We prove the intermediate
claim Luywv :
SNo (u * y + w * v ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (w * v ) Huy Lwv .
We prove the intermediate
claim Luvwy :
SNo (u * v + w * y ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * v ) (w * y ) Huv Lwy .
We prove the intermediate
claim Lwvxy :
SNo (w * v + x * y ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * v ) (x * y ) Lwv Hxy .
We prove the intermediate
claim Lxvwy :
SNo (x * v + w * y ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * v ) (w * y ) Hxv Lwy .
We prove the intermediate
claim Lwvuy :
SNo (w * v + u * y ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * v ) (u * y ) Lwv Huy .
Apply SNoLtE v y Hv Hy Hvy to the current goal.
Let z be given.
Assume Hz2a Hz2b .
We prove the intermediate
claim Lzy :
z ∈ SNoL y .
An
exact proof term for the current goal is
SNoL_I y Hy z Hz1 Hz2b Hz6 .
We prove the intermediate
claim Lzv :
z ∈ SNoR v .
An
exact proof term for the current goal is
SNoR_I v Hv z Hz1 Hz2a Hz5 .
We prove the intermediate
claim Lxz :
SNo (x * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo x z Hx Hz1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu Hz1 .
We prove the intermediate
claim Lwz :
SNo (w * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo w z Hw1 Hz1 .
We prove the intermediate
claim L1 :
w * y + x * z < x * y + w * z .
An exact proof term for the current goal is Hxy1 w Lwx z Lzy .
We prove the intermediate
claim L2 :
w * v + u * z < u * v + w * z .
An exact proof term for the current goal is Huv2 w Lwu z Lzv .
We prove the intermediate
claim L3 :
x * v + w * z < w * v + x * z .
An exact proof term for the current goal is Hxv3 w Lwx z Lzv .
We prove the intermediate
claim L4 :
u * y + w * z < w * y + u * z .
An exact proof term for the current goal is Huy4 w Lwu z Lzy .
We prove the intermediate
claim Lwzwz :
SNo (w * z + w * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * z ) (w * z ) Lwz Lwz .
We will
prove (u * y + x * v ) + (w * z + w * z ) < (x * y + u * v ) + (w * z + w * z ) .
We prove the intermediate
claim Lwyuz :
SNo (w * y + u * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * y ) (u * z ) Lwy Luz .
We prove the intermediate
claim Lwvxz :
SNo (w * v + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * v ) (x * z ) Lwv Lxz .
We prove the intermediate
claim Luywz :
SNo (u * y + w * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (w * z ) Huy Lwz .
We prove the intermediate
claim Lxvwz :
SNo (x * v + w * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * v ) (w * z ) Hxv Lwz .
We prove the intermediate
claim Lwvuz :
SNo (w * v + u * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * v ) (u * z ) Lwv Luz .
We prove the intermediate
claim Lxyxz :
SNo (x * y + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (x * z ) Hxy Lxz .
We prove the intermediate
claim Lwyxz :
SNo (w * y + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (w * y ) (x * z ) Lwy Lxz .
We prove the intermediate
claim Lxywz :
SNo (x * y + w * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (w * z ) Hxy Lwz .
We prove the intermediate
claim Luvwz :
SNo (u * v + w * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * v ) (w * z ) Huv Lwz .
We will
prove (u * y + x * v ) + (w * z + w * z ) < (w * y + u * z ) + (w * v + x * z ) .
We will
prove (u * y + w * z ) + (x * v + w * z ) < (w * y + u * z ) + (w * v + x * z ) .
Apply add_SNo_Lt3 (u * y + w * z ) (x * v + w * z ) (w * y + u * z ) (w * v + x * z ) Luywz Lxvwz Lwyuz Lwvxz to the current goal.
We will
prove u * y + w * z < w * y + u * z .
An exact proof term for the current goal is L4 .
We will
prove x * v + w * z < w * v + x * z .
An exact proof term for the current goal is L3 .
We will
prove (w * y + u * z ) + (w * v + x * z ) < (x * y + u * v ) + (w * z + w * z ) .
We will
prove (w * y + u * z ) + (w * v + x * z ) < (x * y + w * z ) + (u * v + w * z ) .
rewrite the current goal using
add_SNo_com (w * y ) (u * z ) Lwy Luz (from left to right).
rewrite the current goal using
add_SNo_com (w * v ) (u * z ) Lwv Luz (from right to left).
We will
prove (w * v + u * z ) + (w * y + x * z ) < (x * y + w * z ) + (u * v + w * z ) .
rewrite the current goal using
add_SNo_com (w * v + u * z ) (w * y + x * z ) Lwvuz Lwyxz (from left to right).
We will
prove (w * y + x * z ) + (w * v + u * z ) < (x * y + w * z ) + (u * v + w * z ) .
Apply add_SNo_Lt3 (w * y + x * z ) (w * v + u * z ) (x * y + w * z ) (u * v + w * z ) Lwyxz Lwvuz Lxywz Luvwz to the current goal.
An exact proof term for the current goal is L1 .
An exact proof term for the current goal is L2 .
We prove the intermediate
claim Lvy :
v ∈ SNoL y .
An
exact proof term for the current goal is
SNoL_I y Hy v Hv H4 Hvy .
We prove the intermediate
claim L1 :
w * y + x * v < x * y + w * v .
An exact proof term for the current goal is Hxy1 w Lwx v Lvy .
We prove the intermediate
claim L2 :
u * y + w * v < w * y + u * v .
An exact proof term for the current goal is Huy4 w Lwu v Lvy .
We will
prove u * y + x * v < x * y + u * v .
We will
prove w * y + u * y + x * v < w * y + x * y + u * v .
We will
prove w * y + u * y + x * v < u * y + w * v + x * y .
rewrite the current goal using
add_SNo_com_3_0_1 (w * y ) (u * y ) (x * v ) Lwy Huy Hxv (from left to right).
We will
prove u * y + w * y + x * v < u * y + w * v + x * y .
rewrite the current goal using
add_SNo_com (w * v ) (x * y ) Lwv Hxy (from left to right).
An
exact proof term for the current goal is
add_SNo_Lt2 (u * y ) (w * y + x * v ) (x * y + w * v ) Huy Lwyxv Lxywv L1 .
We will
prove u * y + w * v + x * y < w * y + x * y + u * v .
rewrite the current goal using
add_SNo_assoc (u * y ) (w * v ) (x * y ) Huy Lwv Hxy (from left to right).
We will
prove (u * y + w * v ) + x * y < w * y + x * y + u * v .
rewrite the current goal using
add_SNo_com (x * y ) (u * v ) Hxy Huv (from left to right).
We will
prove (u * y + w * v ) + x * y < w * y + u * v + x * y .
rewrite the current goal using
add_SNo_assoc (w * y ) (u * v ) (x * y ) Lwy Huv Hxy (from left to right).
We will
prove (u * y + w * v ) + x * y < (w * y + u * v ) + x * y .
An
exact proof term for the current goal is
add_SNo_Lt1 (u * y + w * v ) (x * y ) (w * y + u * v ) Luywv Hxy Lwyuv L2 .
We prove the intermediate
claim Lyv :
y ∈ SNoR v .
An
exact proof term for the current goal is
SNoR_I v Hv y Hy H4 Hvy .
We prove the intermediate
claim L1 :
x * v + w * y < w * v + x * y .
An exact proof term for the current goal is Hxv3 w Lwx y Lyv .
We prove the intermediate
claim L2 :
w * v + u * y < u * v + w * y .
An exact proof term for the current goal is Huv2 w Lwu y Lyv .
We will
prove u * y + x * v < x * y + u * v .
We will
prove w * y + u * y + x * v < w * y + x * y + u * v .
We will
prove w * y + u * y + x * v < u * y + w * v + x * y .
rewrite the current goal using
add_SNo_rotate_3_1 (u * y ) (x * v ) (w * y ) Huy Hxv Lwy (from right to left).
We will
prove u * y + x * v + w * y < u * y + w * v + x * y .
An
exact proof term for the current goal is
add_SNo_Lt2 (u * y ) (x * v + w * y ) (w * v + x * y ) Huy Lxvwy Lwvxy L1 .
We will
prove u * y + w * v + x * y < w * y + x * y + u * v .
rewrite the current goal using
add_SNo_rotate_3_1 (w * y ) (x * y ) (u * v ) Lwy Hxy Huv (from left to right).
We will
prove u * y + w * v + x * y < u * v + w * y + x * y .
rewrite the current goal using
add_SNo_assoc (u * y ) (w * v ) (x * y ) Huy Lwv Hxy (from left to right).
rewrite the current goal using
add_SNo_assoc (u * v ) (w * y ) (x * y ) Huv Lwy Hxy (from left to right).
We will
prove (u * y + w * v ) + x * y < (u * v + w * y ) + x * y .
rewrite the current goal using
add_SNo_com (u * y ) (w * v ) Huy Lwv (from left to right).
We will
prove (w * v + u * y ) + x * y < (u * v + w * y ) + x * y .
Apply add_SNo_Lt1 (w * v + u * y ) (x * y ) (u * v + w * y ) Lwvuy Hxy Luvwy L2 to the current goal.
We prove the intermediate
claim Lux :
u ∈ SNoL x .
An
exact proof term for the current goal is
SNoL_I x Hx u Hu H1 Hux .
Apply SNoLtE v y Hv Hy Hvy to the current goal.
Let z be given.
Assume Hz2a Hz2b .
We prove the intermediate
claim Lzy :
z ∈ SNoL y .
An
exact proof term for the current goal is
SNoL_I y Hy z Hz1 Hz2b Hz6 .
We prove the intermediate
claim Lzv :
z ∈ SNoR v .
An
exact proof term for the current goal is
SNoR_I v Hv z Hz1 Hz2a Hz5 .
We prove the intermediate
claim Lxz :
SNo (x * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo x z Hx Hz1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu Hz1 .
We prove the intermediate
claim Luyxz :
SNo (u * y + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (x * z ) Huy Lxz .
We prove the intermediate
claim Luvxz :
SNo (u * v + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * v ) (x * z ) Huv Lxz .
We prove the intermediate
claim Lxyuz :
SNo (x * y + u * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (u * z ) Hxy Luz .
We prove the intermediate
claim Lxvuz :
SNo (x * v + u * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * v ) (u * z ) Hxv Luz .
We prove the intermediate
claim L1 :
u * y + x * z < x * y + u * z .
An exact proof term for the current goal is Hxy1 u Lux z Lzy .
We prove the intermediate
claim L2 :
x * v + u * z < u * v + x * z .
An exact proof term for the current goal is Hxv3 u Lux z Lzv .
We will
prove u * y + x * v < x * y + u * v .
We will
prove x * z + u * y + x * v < x * z + x * y + u * v .
We will
prove x * z + u * y + x * v < x * y + u * z + x * v .
rewrite the current goal using
add_SNo_assoc (x * z ) (u * y ) (x * v ) Lxz Huy Hxv (from left to right).
rewrite the current goal using
add_SNo_com (x * z ) (u * y ) Lxz Huy (from left to right).
We will
prove (u * y + x * z ) + x * v < x * y + u * z + x * v .
rewrite the current goal using
add_SNo_assoc (x * y ) (u * z ) (x * v ) Hxy Luz Hxv (from left to right).
An
exact proof term for the current goal is
add_SNo_Lt1 (u * y + x * z ) (x * v ) (x * y + u * z ) Luyxz Hxv Lxyuz L1 .
We will
prove x * y + u * z + x * v < x * z + x * y + u * v .
rewrite the current goal using
add_SNo_rotate_3_1 (x * y ) (u * v ) (x * z ) Hxy Huv Lxz (from right to left).
We will
prove x * y + u * z + x * v < x * y + u * v + x * z .
rewrite the current goal using
add_SNo_com (u * z ) (x * v ) Luz Hxv (from left to right).
We will
prove x * y + x * v + u * z < x * y + u * v + x * z .
An
exact proof term for the current goal is
add_SNo_Lt2 (x * y ) (x * v + u * z ) (u * v + x * z ) Hxy Lxvuz Luvxz L2 .
We prove the intermediate
claim Lvy :
v ∈ SNoL y .
An
exact proof term for the current goal is
SNoL_I y Hy v Hv H4 Hvy .
An exact proof term for the current goal is Hxy1 u Lux v Lvy .
We prove the intermediate
claim Lyv :
y ∈ SNoR v .
An
exact proof term for the current goal is
SNoR_I v Hv y Hy H4 Hvy .
We will
prove u * y + x * v < x * y + u * v .
rewrite the current goal using
add_SNo_com (u * y ) (x * v ) Huy Hxv (from left to right).
rewrite the current goal using
add_SNo_com (x * y ) (u * v ) Hxy Huv (from left to right).
An exact proof term for the current goal is Hxv3 u Lux y Lyv .
We prove the intermediate
claim Lxu :
x ∈ SNoR u .
An
exact proof term for the current goal is
SNoR_I u Hu x Hx H1 Hux .
Apply SNoLtE v y Hv Hy Hvy to the current goal.
Let z be given.
Assume Hz2a Hz2b .
We prove the intermediate
claim Lzy :
z ∈ SNoL y .
An
exact proof term for the current goal is
SNoL_I y Hy z Hz1 Hz2b Hz6 .
We prove the intermediate
claim Lzv :
z ∈ SNoR v .
An
exact proof term for the current goal is
SNoR_I v Hv z Hz1 Hz2a Hz5 .
We prove the intermediate
claim Lxz :
SNo (x * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo x z Hx Hz1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu Hz1 .
We prove the intermediate
claim L1 :
u * y + x * z < x * y + u * z .
An exact proof term for the current goal is Huy4 x Lxu z Lzy .
We prove the intermediate
claim L2 :
x * v + u * z < u * v + x * z .
An exact proof term for the current goal is Huv2 x Lxu z Lzv .
We will
prove u * y + x * v < x * y + u * v .
We will
prove (u * y + x * v ) + x * z < (x * y + u * v ) + x * z .
rewrite the current goal using
add_SNo_com (u * y ) (x * v ) Huy Hxv (from left to right).
We will
prove (x * v + u * y ) + x * z < (x * y + u * v ) + x * z .
rewrite the current goal using
add_SNo_assoc (x * v ) (u * y ) (x * z ) Hxv Huy Lxz (from right to left).
rewrite the current goal using
add_SNo_assoc (x * y ) (u * v ) (x * z ) Hxy Huv Lxz (from right to left).
We will
prove x * v + u * y + x * z < x * y + u * v + x * z .
We prove the intermediate
claim Luyxz :
SNo (u * y + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * y ) (x * z ) Huy Lxz .
We prove the intermediate
claim Luzxy :
SNo (u * z + x * y ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * z ) (x * y ) Luz Hxy .
We prove the intermediate
claim Luvxz :
SNo (u * v + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (u * v ) (x * z ) Huv Lxz .
We will
prove x * v + u * y + x * z < x * v + u * z + x * y .
rewrite the current goal using
add_SNo_com (u * z ) (x * y ) Luz Hxy (from left to right).
We will
prove x * v + u * z + x * y < x * y + u * v + x * z .
rewrite the current goal using
add_SNo_rotate_3_1 (x * v ) (u * z ) (x * y ) Hxv Luz Hxy (from left to right).
We will
prove x * y + x * v + u * z < x * y + u * v + x * z .
We prove the intermediate
claim Lvy :
v ∈ SNoL y .
An
exact proof term for the current goal is
SNoL_I y Hy v Hv H4 Hvy .
We will
prove u * y + x * v < x * y + u * v .
An exact proof term for the current goal is Huy4 x Lxu v Lvy .
We prove the intermediate
claim Lyv :
y ∈ SNoR v .
An
exact proof term for the current goal is
SNoR_I v Hv y Hy H4 Hvy .
We will
prove u * y + x * v < x * y + u * v .
rewrite the current goal using
add_SNo_com (u * y ) (x * v ) Huy Hxv (from left to right).
rewrite the current goal using
add_SNo_com (x * y ) (u * v ) Hxy Huv (from left to right).
An exact proof term for the current goal is Huv2 x Lxu y Lyv .
∎
Proof: Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy .
Apply SNoLeE u x Hu Hx Hux to the current goal.
Apply SNoLeE v y Hv Hy Hvy to the current goal.
An
exact proof term for the current goal is
mul_SNo_Lt x y u v Hx Hy Hu Hv Hux Hvy .
rewrite the current goal using Hvy (from left to right).
We will
prove u * y + x * y ≤ x * y + u * y .
We will
prove x * y + u * y ≤ x * y + u * y .
rewrite the current goal using Hux (from left to right).
We will
prove x * y + x * v ≤ x * y + x * v .
∎
Proof: Let x and y be given.
Assume Hx Hy .
Set P to be the term
λu ⇒ P1 u ∨ P2 u of type
set → prop .
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 .
Let u be given.
Apply dneg to the current goal.
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
Apply SNoLtLe_tra u (x * y ) u Hu1 Lxy Hu1 Hu3 to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 .
rewrite the current goal using HE (from left to right).
rewrite the current goal using
SNo_eta u Hu1 (from left to right).
Let z be given.
Assume Hz .
rewrite the current goal using
SNo_eta u Hu1 (from right to left).
Apply HLE z Hz to the current goal.
Let v be given.
Let w be given.
rewrite the current goal using Hze (from left to right).
We will
prove v * y + x * w + - v * w < u .
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We will
prove v * y + x * w < u + v * w .
An exact proof term for the current goal is L1 v Hv w Hw .
Let v be given.
Let w be given.
rewrite the current goal using Hze (from left to right).
We will
prove v * y + x * w + - v * w < u .
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We will
prove v * y + x * w < u + v * w .
An exact proof term for the current goal is L2 v Hv w Hw .
Let z be given.
rewrite the current goal using HE (from right to left).
Apply SNoR_E u Hu1 z Hz to the current goal.
We prove the intermediate claim LPz : P z .
Apply IH z to the current goal.
An exact proof term for the current goal is Hz1 .
An exact proof term for the current goal is Hu1 .
An exact proof term for the current goal is Hz2 .
An exact proof term for the current goal is H1 .
Apply LPz to the current goal.
Assume H2 : P1 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim L3 :
z + v * w < u + v * w .
We will
prove v * y + x * w < u + v * w .
An exact proof term for the current goal is L1 v Hv w Hw .
We prove the intermediate
claim L4 :
z < u .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L4 .
Assume H2 : P2 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim L5 :
z + v * w < u + v * w .
We will
prove v * y + x * w < u + v * w .
An exact proof term for the current goal is L2 v Hv w Hw .
We prove the intermediate
claim L6 :
z < u .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L6 .
rewrite the current goal using H1 (from right to left).
An exact proof term for the current goal is Hz2 .
An exact proof term for the current goal is H1 .
Let u be given.
Apply SNoL_E (x * y ) Lxy u Hu to the current goal.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3 .
∎
Proof: Let x and y be given.
Assume Hx Hy .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp1 Hp2 .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw .
∎
Proof: Let x and y be given.
Assume Hx Hy .
Set P to be the term
λu ⇒ P1 u ∨ P2 u of type
set → prop .
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 .
Let u be given.
Apply dneg to the current goal.
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
Apply SNoLtLe_tra (x * y ) u (x * y ) Lxy Hu1 Lxy Hu3 to the current goal.
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 .
rewrite the current goal using HE (from left to right).
rewrite the current goal using
SNo_eta u Hu1 (from left to right).
Let z be given.
rewrite the current goal using HE (from right to left).
Apply SNoL_E u Hu1 z Hz to the current goal.
An exact proof term for the current goal is H1 .
rewrite the current goal using H1 (from right to left).
An exact proof term for the current goal is Hz2 .
We prove the intermediate claim LPz : P z .
Apply IH z to the current goal.
An exact proof term for the current goal is Hz1 .
An exact proof term for the current goal is Hu1 .
An exact proof term for the current goal is Hz2 .
An exact proof term for the current goal is H1 .
Apply LPz to the current goal.
Assume H2 : P1 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim L3 :
u + v * w < z + v * w .
We will
prove u + v * w < v * y + x * w .
An exact proof term for the current goal is L1 v Hv w Hw .
An exact proof term for the current goal is Hvw .
We prove the intermediate
claim L4 :
u < z .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 L4 Hz3 .
Assume H2 : P2 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim L5 :
u + v * w < z + v * w .
We will
prove u + v * w < v * y + x * w .
An exact proof term for the current goal is L2 v Hv w Hw .
An exact proof term for the current goal is Hvw .
We prove the intermediate
claim L6 :
u < z .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 L6 Hz3 .
Let z be given.
Assume Hz .
rewrite the current goal using
SNo_eta u Hu1 (from right to left).
Apply HRE z Hz to the current goal.
Let v be given.
Let w be given.
rewrite the current goal using Hze (from left to right).
We will
prove u < v * y + x * w + - v * w .
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We will
prove u + v * w < v * y + x * w .
An exact proof term for the current goal is L1 v Hv w Hw .
Let v be given.
Let w be given.
rewrite the current goal using Hze (from left to right).
We will
prove u < v * y + x * w + - v * w .
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We will
prove u + v * w < v * y + x * w .
An exact proof term for the current goal is L2 v Hv w Hw .
Let u be given.
Apply SNoR_E (x * y ) Lxy u Hu to the current goal.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3 .
∎
Proof: Let x and y be given.
Assume Hx Hy .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp1 Hp2 .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw .
∎
Theorem. (
mul_SNo_Subq_lem ) The following is provable:
∀x y X Y Z W, ∀U U', (∀u, u ∈ U → (∀q : prop , (∀ w0 ∈ X , ∀ w1 ∈ Y , u = w0 * y + x * w1 + - w0 * w1 → q ) → (∀ z0 ∈ Z , ∀ z1 ∈ W , u = z0 * y + x * z1 + - z0 * z1 → q ) → q ) ) → (∀ w0 ∈ X , ∀ w1 ∈ Y , w0 * y + x * w1 + - w0 * w1 ∈ U' ) → (∀ w0 ∈ Z , ∀ w1 ∈ W , w0 * y + x * w1 + - w0 * w1 ∈ U' ) → U ⊆ U'
Proof: Let x, y, X, Y, Z, W, U and U' be given.
Assume HUE HU'I1 HU'I2 .
Let u be given.
Apply HUE u Hu to the current goal.
Let w be given.
Assume Hw .
Let z be given.
Assume Hz Huwz .
rewrite the current goal using Huwz (from left to right).
Apply HU'I1 to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is Hz .
Let w be given.
Assume Hw .
Let z be given.
Assume Hz Huwz .
rewrite the current goal using Huwz (from left to right).
Apply HU'I2 to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is Hz .
∎
Proof: Let x be given.
Let L and R be given.
Assume HLE HLI1 HLI2 HRE HRI1 HRI2 Hx0 .
rewrite the current goal using Hx0 (from left to right).
We prove the intermediate
claim LL0 :
L = 0 .
Let w be given.
Apply HLE w Hw to the current goal.
Let u be given.
Let v be given.
rewrite the current goal using
SNoL_0 (from left to right).
An
exact proof term for the current goal is
EmptyE v Hv .
Let u be given.
Let v be given.
rewrite the current goal using
SNoR_0 (from left to right).
An
exact proof term for the current goal is
EmptyE v Hv .
We prove the intermediate
claim LR0 :
R = 0 .
Let w be given.
Apply HRE w Hw to the current goal.
Let u be given.
Let v be given.
rewrite the current goal using
SNoR_0 (from left to right).
An
exact proof term for the current goal is
EmptyE v Hv .
Let u be given.
Let v be given.
rewrite the current goal using
SNoL_0 (from left to right).
An
exact proof term for the current goal is
EmptyE v Hv .
rewrite the current goal using LL0 (from left to right).
rewrite the current goal using LR0 (from left to right).
An
exact proof term for the current goal is
SNoCut_0_0 .
∎
Proof:
Let x be given.
Assume Hx .
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 Hx1e .
We prove the intermediate
claim L0L1 :
0 ∈ SNoL 1 .
rewrite the current goal using
SNoL_1 (from left to right).
An
exact proof term for the current goal is
In_0_1 .
rewrite the current goal using Hx1e (from left to right).
rewrite the current goal using
SNo_eta x Hx (from left to right).
Let w be given.
Assume Hw .
We will
prove w ∈ SNoL x .
Apply HLE w Hw to the current goal.
Let u be given.
Let v be given.
rewrite the current goal using
SNoL_1 (from left to right).
Apply cases_1 v Hv to the current goal.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _ .
We prove the intermediate
claim L1 :
w = u .
rewrite the current goal using Hwuv (from left to right).
Use transitivity with
u * 1 + 0 , and
u * 1 .
Use f_equal.
Use transitivity with
x * 0 + 0 , and
x * 0 .
Use f_equal.
We will
prove - u * 0 = 0 .
Use transitivity with and
- 0 .
An
exact proof term for the current goal is
minus_SNo_0 .
An
exact proof term for the current goal is
IHx u (SNoL_SNoS x Hx u Hu ) .
rewrite the current goal using L1 (from left to right).
An exact proof term for the current goal is Hu .
Let u be given.
Let v be given.
rewrite the current goal using
SNoR_1 (from left to right).
An
exact proof term for the current goal is
EmptyE v Hv .
Let z be given.
Assume Hz .
We will
prove z ∈ SNoR x .
Apply HRE z Hz to the current goal.
Let u be given.
Let v be given.
rewrite the current goal using
SNoR_1 (from left to right).
An
exact proof term for the current goal is
EmptyE v Hv .
Let u be given.
Let v be given.
rewrite the current goal using
SNoL_1 (from left to right).
Apply cases_1 v Hv to the current goal.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _ .
We prove the intermediate
claim L1 :
z = u .
rewrite the current goal using Hzuv (from left to right).
Use transitivity with
u * 1 + 0 , and
u * 1 .
Use f_equal.
Use transitivity with
x * 0 + 0 , and
x * 0 .
Use f_equal.
We will
prove - u * 0 = 0 .
Use transitivity with and
- 0 .
An
exact proof term for the current goal is
minus_SNo_0 .
An
exact proof term for the current goal is
IHx u (SNoR_SNoS x Hx u Hu ) .
rewrite the current goal using L1 (from left to right).
An exact proof term for the current goal is Hu .
Let w be given.
Assume Hw .
Apply SNoL_E x Hx w Hw to the current goal.
Assume Hwa _ _ .
rewrite the current goal using Hx1e (from right to left).
We prove the intermediate
claim L1 :
w * 1 + x * 0 = w .
Use transitivity with
w * 1 + 0 , and
w * 1 .
An
exact proof term for the current goal is
IHx w (SNoL_SNoS x Hx w Hw ) .
We prove the intermediate
claim L2 :
x * 1 + w * 0 = x * 1 .
Use transitivity with and
x * 1 + 0 .
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
An
exact proof term for the current goal is
Hx11 w Hw 0 L0L1 .
Let z be given.
Assume Hz .
Apply SNoR_E x Hx z Hz to the current goal.
Assume Hza _ _ .
rewrite the current goal using Hx1e (from right to left).
We prove the intermediate
claim L1 :
x * 1 + z * 0 = x * 1 .
Use transitivity with and
x * 1 + 0 .
We prove the intermediate
claim L2 :
z * 1 + x * 0 = z .
Use transitivity with
z * 1 + 0 , and
z * 1 .
An
exact proof term for the current goal is
IHx z (SNoR_SNoS x Hx z Hz ) .
rewrite the current goal using L2 (from right to left).
rewrite the current goal using L1 (from right to left).
An
exact proof term for the current goal is
Hx14 z Hz 0 L0L1 .
∎
Proof:
Let x and y be given.
Assume Hx Hy .
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 Hxye .
Let L' and R' be given.
Assume HL'R' HL'E HL'I1 HL'I2 HR'E HR'I1 HR'I2 Hyxe .
rewrite the current goal using Hxye (from left to right).
rewrite the current goal using Hyxe (from left to right).
We prove the intermediate
claim LLL' :
L = L' .
Let u be given.
Let v be given.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoL_SNoS x Hx u Hu ) (from left to right).
rewrite the current goal using
IHy v (SNoL_SNoS y Hy v Hv ) (from left to right).
rewrite the current goal using
IHxy u (SNoL_SNoS x Hx u Hu ) v (SNoL_SNoS y Hy v Hv ) (from left to right).
Apply HL'I1 to the current goal.
An exact proof term for the current goal is Hv .
An exact proof term for the current goal is Hu .
Let u be given.
Let v be given.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoR_SNoS x Hx u Hu ) (from left to right).
rewrite the current goal using
IHy v (SNoR_SNoS y Hy v Hv ) (from left to right).
rewrite the current goal using
IHxy u (SNoR_SNoS x Hx u Hu ) v (SNoR_SNoS y Hy v Hv ) (from left to right).
Apply HL'I2 to the current goal.
An exact proof term for the current goal is Hv .
An exact proof term for the current goal is Hu .
Let v be given.
Let u be given.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoL_SNoS x Hx u Hu ) (from right to left).
rewrite the current goal using
IHy v (SNoL_SNoS y Hy v Hv ) (from right to left).
rewrite the current goal using
IHxy u (SNoL_SNoS x Hx u Hu ) v (SNoL_SNoS y Hy v Hv ) (from right to left).
Apply HLI1 to the current goal.
An exact proof term for the current goal is Hu .
An exact proof term for the current goal is Hv .
Let v be given.
Let u be given.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoR_SNoS x Hx u Hu ) (from right to left).
rewrite the current goal using
IHy v (SNoR_SNoS y Hy v Hv ) (from right to left).
rewrite the current goal using
IHxy u (SNoR_SNoS x Hx u Hu ) v (SNoR_SNoS y Hy v Hv ) (from right to left).
Apply HLI2 to the current goal.
An exact proof term for the current goal is Hu .
An exact proof term for the current goal is Hv .
We prove the intermediate
claim LRR' :
R = R' .
Let u be given.
Let v be given.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoL_SNoS x Hx u Hu ) (from left to right).
rewrite the current goal using
IHy v (SNoR_SNoS y Hy v Hv ) (from left to right).
rewrite the current goal using
IHxy u (SNoL_SNoS x Hx u Hu ) v (SNoR_SNoS y Hy v Hv ) (from left to right).
Apply HR'I2 to the current goal.
An exact proof term for the current goal is Hv .
An exact proof term for the current goal is Hu .
Let u be given.
Let v be given.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoR_SNoS x Hx u Hu ) (from left to right).
rewrite the current goal using
IHy v (SNoL_SNoS y Hy v Hv ) (from left to right).
rewrite the current goal using
IHxy u (SNoR_SNoS x Hx u Hu ) v (SNoL_SNoS y Hy v Hv ) (from left to right).
Apply HR'I1 to the current goal.
An exact proof term for the current goal is Hv .
An exact proof term for the current goal is Hu .
Let v be given.
Let u be given.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoR_SNoS x Hx u Hu ) (from right to left).
rewrite the current goal using
IHy v (SNoL_SNoS y Hy v Hv ) (from right to left).
rewrite the current goal using
IHxy u (SNoR_SNoS x Hx u Hu ) v (SNoL_SNoS y Hy v Hv ) (from right to left).
Apply HRI2 to the current goal.
An exact proof term for the current goal is Hu .
An exact proof term for the current goal is Hv .
Let v be given.
Let u be given.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hua _ _ .
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
rewrite the current goal using
IHx u (SNoL_SNoS x Hx u Hu ) (from right to left).
rewrite the current goal using
IHy v (SNoR_SNoS y Hy v Hv ) (from right to left).
rewrite the current goal using
IHxy u (SNoL_SNoS x Hx u Hu ) v (SNoR_SNoS y Hy v Hv ) (from right to left).
Apply HRI1 to the current goal.
An exact proof term for the current goal is Hu .
An exact proof term for the current goal is Hv .
rewrite the current goal using LLL' (from left to right).
rewrite the current goal using LRR' (from left to right).
Use reflexivity.
∎
Proof:
Let x and y be given.
Assume Hx Hy .
We prove the intermediate
claim Lmx :
SNo (- x ) .
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 .
Let L' and R' be given.
Assume HL'R' HL'E HL'I1 HL'I2 HR'E HR'I1 HR'I2 .
rewrite the current goal using Hxye (from left to right).
rewrite the current goal using L1 (from left to right).
rewrite the current goal using Hmxye (from left to right).
Use f_equal.
We will
prove L' = { - z | z ∈ R } .
Let u be given.
Let v be given.
Apply SNoL_E (- x ) Lmx u Hu to the current goal.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoR x .
Apply SNoR_I x Hx (- u ) Lmu1 to the current goal.
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
u * y + (- x ) * v + - u * v = - ((- u ) * y + x * v + - (- u ) * v ) .
Use symmetry.
Use transitivity with and
- (- u ) * y + - x * v + - - (- u ) * v .
Use f_equal.
We will
prove - (- u ) * y = u * y .
Use transitivity with and
(- - u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx (- u ) (SNoR_SNoS x Hx (- u ) Lmu2 ) .
Use f_equal.
We will
prove - x * v = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoL_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (- u ) * v = u * v .
Use transitivity with and
(- - u ) * v .
Use symmetry.
We will
prove (- - u ) * v = - (- u ) * v .
An
exact proof term for the current goal is
IHxy (- u ) (SNoR_SNoS x Hx (- u ) Lmu2 ) v (SNoL_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
Apply ReplI R (λz ⇒ - z ) to the current goal.
We will
prove (- u ) * y + x * v + - (- u ) * v ∈ R .
Apply HRI2 to the current goal.
An exact proof term for the current goal is Lmu2 .
We will
prove v ∈ SNoL y .
An exact proof term for the current goal is Hv .
Let u be given.
Let v be given.
Apply SNoR_E (- x ) Lmx u Hu to the current goal.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoL x .
Apply SNoL_I x Hx (- u ) Lmu1 to the current goal.
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
u * y + (- x ) * v + - u * v = - ((- u ) * y + x * v + - (- u ) * v ) .
Use symmetry.
Use transitivity with and
- (- u ) * y + - x * v + - - (- u ) * v .
Use f_equal.
We will
prove - (- u ) * y = u * y .
Use transitivity with and
(- - u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx (- u ) (SNoL_SNoS x Hx (- u ) Lmu2 ) .
Use f_equal.
We will
prove - x * v = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoR_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (- u ) * v = u * v .
Use transitivity with and
(- - u ) * v .
Use symmetry.
We will
prove (- - u ) * v = - (- u ) * v .
An
exact proof term for the current goal is
IHxy (- u ) (SNoL_SNoS x Hx (- u ) Lmu2 ) v (SNoR_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
Apply ReplI R (λz ⇒ - z ) to the current goal.
We will
prove (- u ) * y + x * v + - (- u ) * v ∈ R .
Apply HRI1 to the current goal.
An exact proof term for the current goal is Lmu2 .
We will
prove v ∈ SNoR y .
An exact proof term for the current goal is Hv .
We will
prove { - z | z ∈ R } ⊆ L' .
Let a be given.
Assume Ha .
Let z be given.
rewrite the current goal using Hze (from left to right).
Apply HRE z Hz to the current goal.
Let u be given.
Let v be given.
Apply SNoL_E x Hx u Hu to the current goal.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoR (- x ) .
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
- z = (- u ) * y + (- x ) * v + - (- u ) * v .
rewrite the current goal using Hzuv (from left to right).
Use f_equal.
We will
prove - (u * y ) = (- u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx u (SNoL_SNoS x Hx u Hu ) .
Use f_equal.
We will
prove - (x * v ) = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoR_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (u * v ) = (- u ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHxy u (SNoL_SNoS x Hx u Hu ) v (SNoR_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
We will
prove (- u ) * y + (- x ) * v + - (- u ) * v ∈ L' .
Apply HL'I2 to the current goal.
An exact proof term for the current goal is Lmu2 .
An exact proof term for the current goal is Hv .
Let u be given.
Let v be given.
Apply SNoR_E x Hx u Hu to the current goal.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoL (- x ) .
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
- z = (- u ) * y + (- x ) * v + - (- u ) * v .
rewrite the current goal using Hzuv (from left to right).
Use f_equal.
We will
prove - (u * y ) = (- u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx u (SNoR_SNoS x Hx u Hu ) .
Use f_equal.
We will
prove - (x * v ) = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoL_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (u * v ) = (- u ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHxy u (SNoR_SNoS x Hx u Hu ) v (SNoL_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
We will
prove (- u ) * y + (- x ) * v + - (- u ) * v ∈ L' .
Apply HL'I1 to the current goal.
An exact proof term for the current goal is Lmu2 .
An exact proof term for the current goal is Hv .
We will
prove R' = { - w | w ∈ L } .
Let u be given.
Let v be given.
Apply SNoL_E (- x ) Lmx u Hu to the current goal.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoR x .
Apply SNoR_I x Hx (- u ) Lmu1 to the current goal.
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
u * y + (- x ) * v + - u * v = - ((- u ) * y + x * v + - (- u ) * v ) .
Use symmetry.
Use transitivity with and
- (- u ) * y + - x * v + - - (- u ) * v .
Use f_equal.
We will
prove - (- u ) * y = u * y .
Use transitivity with and
(- - u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx (- u ) (SNoR_SNoS x Hx (- u ) Lmu2 ) .
Use f_equal.
We will
prove - x * v = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoR_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (- u ) * v = u * v .
Use transitivity with and
(- - u ) * v .
Use symmetry.
We will
prove (- - u ) * v = - (- u ) * v .
An
exact proof term for the current goal is
IHxy (- u ) (SNoR_SNoS x Hx (- u ) Lmu2 ) v (SNoR_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
Apply ReplI L (λw ⇒ - w ) to the current goal.
We will
prove (- u ) * y + x * v + - (- u ) * v ∈ L .
Apply HLI2 to the current goal.
An exact proof term for the current goal is Lmu2 .
We will
prove v ∈ SNoR y .
An exact proof term for the current goal is Hv .
Let u be given.
Let v be given.
Apply SNoR_E (- x ) Lmx u Hu to the current goal.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoL x .
Apply SNoL_I x Hx (- u ) Lmu1 to the current goal.
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from right to left).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
u * y + (- x ) * v + - u * v = - ((- u ) * y + x * v + - (- u ) * v ) .
Use symmetry.
Use transitivity with and
- (- u ) * y + - x * v + - - (- u ) * v .
Use f_equal.
We will
prove - (- u ) * y = u * y .
Use transitivity with and
(- - u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx (- u ) (SNoL_SNoS x Hx (- u ) Lmu2 ) .
Use f_equal.
We will
prove - x * v = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoL_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (- u ) * v = u * v .
Use transitivity with and
(- - u ) * v .
Use symmetry.
We will
prove (- - u ) * v = - (- u ) * v .
An
exact proof term for the current goal is
IHxy (- u ) (SNoL_SNoS x Hx (- u ) Lmu2 ) v (SNoL_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
Apply ReplI L (λw ⇒ - w ) to the current goal.
We will
prove (- u ) * y + x * v + - (- u ) * v ∈ L .
Apply HLI1 to the current goal.
An exact proof term for the current goal is Lmu2 .
We will
prove v ∈ SNoL y .
An exact proof term for the current goal is Hv .
We will
prove { - w | w ∈ L } ⊆ R' .
Let a be given.
Assume Ha .
Let w be given.
rewrite the current goal using Hwe (from left to right).
Apply HLE w Hw to the current goal.
Let u be given.
Let v be given.
Apply SNoL_E x Hx u Hu to the current goal.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoR (- x ) .
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
- w = (- u ) * y + (- x ) * v + - (- u ) * v .
rewrite the current goal using Hwuv (from left to right).
Use f_equal.
We will
prove - (u * y ) = (- u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx u (SNoL_SNoS x Hx u Hu ) .
Use f_equal.
We will
prove - (x * v ) = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoL_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (u * v ) = (- u ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHxy u (SNoL_SNoS x Hx u Hu ) v (SNoL_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
We will
prove (- u ) * y + (- x ) * v + - (- u ) * v ∈ R' .
Apply HR'I2 to the current goal.
An exact proof term for the current goal is Lmu2 .
An exact proof term for the current goal is Hv .
Let u be given.
Let v be given.
Apply SNoR_E x Hx u Hu to the current goal.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hva _ _ .
We prove the intermediate
claim Lmu1 :
SNo (- u ) .
We prove the intermediate
claim Lmu2 :
- u ∈ SNoL (- x ) .
rewrite the current goal using
minus_SNo_Lev u Hua (from left to right).
rewrite the current goal using
minus_SNo_Lev x Hx (from left to right).
An exact proof term for the current goal is Hub .
We prove the intermediate
claim L1 :
- w = (- u ) * y + (- x ) * v + - (- u ) * v .
rewrite the current goal using Hwuv (from left to right).
Use f_equal.
We will
prove - (u * y ) = (- u ) * y .
Use symmetry.
An
exact proof term for the current goal is
IHx u (SNoR_SNoS x Hx u Hu ) .
Use f_equal.
We will
prove - (x * v ) = (- x ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHy v (SNoR_SNoS y Hy v Hv ) .
Use f_equal.
We will
prove - (u * v ) = (- u ) * v .
Use symmetry.
An
exact proof term for the current goal is
IHxy u (SNoR_SNoS x Hx u Hu ) v (SNoR_SNoS y Hy v Hv ) .
rewrite the current goal using L1 (from left to right).
We will
prove (- u ) * y + (- x ) * v + - (- u ) * v ∈ R' .
Apply HR'I1 to the current goal.
An exact proof term for the current goal is Lmu2 .
An exact proof term for the current goal is Hv .
∎
Proof: Let x and y be given.
Assume Hx Hy .
rewrite the current goal using
mul_SNo_com x y Hx Hy (from left to right).
∎
Proof: Set P to be the term
λx y z ⇒ (x + y ) * z = x * z + y * z of type
set → set → set → prop .
We will
prove ∀x y z, SNo x → SNo y → SNo z → P x y z .
Let x, y and z be given.
Assume Hx Hy Hz .
We will
prove (x + y ) * z = x * z + y * z .
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 .
We prove the intermediate
claim Lxy :
SNo (x + y ) .
An
exact proof term for the current goal is
SNo_add_SNo x y Hx Hy .
We prove the intermediate
claim Lxyz :
SNo ((x + y ) * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) z Lxy Hz .
We prove the intermediate
claim Lxz :
SNo (x * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo x z Hx Hz .
We prove the intermediate
claim Lyz :
SNo (y * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo y z Hy Hz .
We prove the intermediate
claim Lxzyz :
SNo (x * z + y * z ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * z ) (y * z ) Lxz Lyz .
We prove the intermediate
claim LE :
x * z + y * z = SNoCut (L1 ∪ L2 ) (R1 ∪ R2 ) .
rewrite the current goal using HE (from left to right).
rewrite the current goal using LE (from left to right).
An exact proof term for the current goal is HLR .
Let u be given.
rewrite the current goal using LE (from right to left).
We will
prove u < x * z + y * z .
Apply HLE u Hu to the current goal.
Let v be given.
Let w be given.
rewrite the current goal using Hue (from left to right).
We will
prove v * z + (x + y ) * w + - v * w < x * z + y * z .
Apply SNoL_E (x + y ) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lxwyw :
SNo (x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * w ) (y * w ) Lxw Lyw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvzxwyw :
SNo (v * z + x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (v * z ) (x * w ) (y * w ) Lvz Lxw Lyw .
We prove the intermediate
claim Lxzyzvw :
SNo (x * z + y * z + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw .
We will
prove v * z + (x + y ) * w < (x * z + y * z ) + v * w .
rewrite the current goal using
IHz w (SNoL_SNoS z Hz w Hw ) (from left to right).
We will
prove v * z + x * w + y * w < (x * z + y * z ) + v * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw (from right to left).
We will
prove v * z + x * w + y * w < x * z + y * z + v * w .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Luy :
SNo (u + y ) .
An
exact proof term for the current goal is
SNo_add_SNo u y Hu1 Hy .
We will
prove (v * z + x * w + y * w ) + u * w < (x * z + y * z + v * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w .
We will
prove v * z + x * w + y * w + u * w ≤ u * z + y * z + v * w + x * w .
We will
prove x * w + v * z + y * w + u * w ≤ u * z + y * z + v * w + x * w .
rewrite the current goal using
add_SNo_rotate_4_1 (u * z ) (y * z ) (v * w ) (x * w ) Luz Lyz Lvw Lxw (from left to right).
We will
prove x * w + v * z + y * w + u * w ≤ x * w + u * z + y * z + v * w .
We will
prove v * z + y * w + u * w ≤ u * z + y * z + v * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove v * z + u * w + y * w ≤ u * z + y * z + v * w .
rewrite the current goal using
IHxz u (SNoL_SNoS x Hx u Hu ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + (u + y ) * w ≤ u * z + y * z + v * w .
rewrite the current goal using
add_SNo_assoc (u * z ) (y * z ) (v * w ) Luz Lyz Lvw (from left to right).
We will
prove v * z + (u + y ) * w ≤ (u * z + y * z ) + v * w .
rewrite the current goal using
IHx u (SNoL_SNoS x Hx u Hu ) (from right to left).
We will
prove v * z + (u + y ) * w ≤ (u + y ) * z + v * w .
Apply mul_SNo_Le (u + y ) z v w Luy Hz Hv1 Hw1 to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
We will
prove u * z + y * z + v * w + x * w < x * z + y * z + v * w + u * w .
We will
prove y * z + u * z + v * w + x * w < y * z + x * z + v * w + u * w .
We will
prove u * z + v * w + x * w < x * z + v * w + u * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (x * w ) Luz Lvw Lxw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (x * z ) (v * w ) (u * w ) Lxz Lvw Luw (from left to right).
We will
prove v * w + u * z + x * w < v * w + x * z + u * w .
We will
prove u * z + x * w < x * z + u * w .
An
exact proof term for the current goal is
mul_SNo_Lt x z u w Hx Hz Hu1 Hw1 Hu3 Hw3 .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoL_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Lxu :
SNo (x + u ) .
An
exact proof term for the current goal is
SNo_add_SNo x u Hx Hu1 .
We will
prove (v * z + x * w + y * w ) + u * w < (x * z + y * z + v * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w .
We will
prove v * z + x * w + y * w + u * w ≤ x * z + u * z + v * w + y * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove v * z + x * w + u * w + y * w ≤ x * z + u * z + v * w + y * w .
rewrite the current goal using
add_SNo_rotate_4_1 (v * z ) (x * w ) (u * w ) (y * w ) Lvz Lxw Luw Lyw (from left to right).
We will
prove y * w + v * z + x * w + u * w ≤ x * z + u * z + v * w + y * w .
rewrite the current goal using
add_SNo_rotate_4_1 (x * z ) (u * z ) (v * w ) (y * w ) Lxz Luz Lvw Lyw (from left to right).
We will
prove y * w + v * z + x * w + u * w ≤ y * w + x * z + u * z + v * w .
We will
prove v * z + x * w + u * w ≤ x * z + u * z + v * w .
rewrite the current goal using
IHyz u (SNoL_SNoS y Hy u Hu ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + (x + u ) * w ≤ x * z + u * z + v * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (u * z ) (v * w ) Lxz Luz Lvw (from left to right).
We will
prove v * z + (x + u ) * w ≤ (x * z + u * z ) + v * w .
rewrite the current goal using
IHy u (SNoL_SNoS y Hy u Hu ) (from right to left).
We will
prove v * z + (x + u ) * w ≤ (x + u ) * z + v * w .
Apply mul_SNo_Le (x + u ) z v w Lxu Hz Hv1 Hw1 to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
We will
prove x * z + u * z + v * w + y * w < x * z + y * z + v * w + u * w .
We will
prove u * z + v * w + y * w < y * z + v * w + u * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (y * w ) Luz Lvw Lyw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (y * z ) (v * w ) (u * w ) Lyz Lvw Luw (from left to right).
We will
prove v * w + u * z + y * w < v * w + y * z + u * w .
We will
prove u * z + y * w < y * z + u * w .
An
exact proof term for the current goal is
mul_SNo_Lt y z u w Hy Hz Hu1 Hw1 Hu3 Hw3 .
Let v be given.
Let w be given.
rewrite the current goal using Hue (from left to right).
We will
prove v * z + (x + y ) * w + - v * w < x * z + y * z .
Apply SNoR_E (x + y ) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lxwyw :
SNo (x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * w ) (y * w ) Lxw Lyw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvzxwyw :
SNo (v * z + x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (v * z ) (x * w ) (y * w ) Lvz Lxw Lyw .
We prove the intermediate
claim Lxzyzvw :
SNo (x * z + y * z + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw .
We will
prove v * z + (x + y ) * w < (x * z + y * z ) + v * w .
rewrite the current goal using
IHz w (SNoR_SNoS z Hz w Hw ) (from left to right).
We will
prove v * z + x * w + y * w < (x * z + y * z ) + v * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw (from right to left).
We will
prove v * z + x * w + y * w < x * z + y * z + v * w .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Luy :
SNo (u + y ) .
An
exact proof term for the current goal is
SNo_add_SNo u y Hu1 Hy .
We will
prove (v * z + x * w + y * w ) + u * w < (x * z + y * z + v * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w .
We will
prove v * z + x * w + y * w + u * w ≤ u * z + y * z + v * w + x * w .
We will
prove x * w + v * z + y * w + u * w ≤ u * z + y * z + v * w + x * w .
rewrite the current goal using
add_SNo_rotate_4_1 (u * z ) (y * z ) (v * w ) (x * w ) Luz Lyz Lvw Lxw (from left to right).
We will
prove x * w + v * z + y * w + u * w ≤ x * w + u * z + y * z + v * w .
We will
prove v * z + y * w + u * w ≤ u * z + y * z + v * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove v * z + u * w + y * w ≤ u * z + y * z + v * w .
rewrite the current goal using
IHxz u (SNoR_SNoS x Hx u Hu ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + (u + y ) * w ≤ u * z + y * z + v * w .
rewrite the current goal using
add_SNo_assoc (u * z ) (y * z ) (v * w ) Luz Lyz Lvw (from left to right).
We will
prove v * z + (u + y ) * w ≤ (u * z + y * z ) + v * w .
rewrite the current goal using
IHx u (SNoR_SNoS x Hx u Hu ) (from right to left).
We will
prove v * z + (u + y ) * w ≤ (u + y ) * z + v * w .
Apply mul_SNo_Le v w (u + y ) z Hv1 Hw1 Luy Hz to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
We will
prove u * z + y * z + v * w + x * w < x * z + y * z + v * w + u * w .
We will
prove y * z + u * z + v * w + x * w < y * z + x * z + v * w + u * w .
We will
prove u * z + v * w + x * w < x * z + v * w + u * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (x * w ) Luz Lvw Lxw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (x * z ) (v * w ) (u * w ) Lxz Lvw Luw (from left to right).
We will
prove v * w + u * z + x * w < v * w + x * z + u * w .
We will
prove u * z + x * w < x * z + u * w .
rewrite the current goal using
add_SNo_com (x * z ) (u * w ) Lxz Luw (from left to right).
rewrite the current goal using
add_SNo_com (u * z ) (x * w ) Luz Lxw (from left to right).
We will
prove x * w + u * z < u * w + x * z .
An
exact proof term for the current goal is
mul_SNo_Lt u w x z Hu1 Hw1 Hx Hz Hu3 Hw3 .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoR_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Lxu :
SNo (x + u ) .
An
exact proof term for the current goal is
SNo_add_SNo x u Hx Hu1 .
We will
prove (v * z + x * w + y * w ) + u * w < (x * z + y * z + v * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove v * z + x * w + y * w + u * w < x * z + y * z + v * w + u * w .
We will
prove v * z + x * w + y * w + u * w ≤ x * z + u * z + v * w + y * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove v * z + x * w + u * w + y * w ≤ x * z + u * z + v * w + y * w .
rewrite the current goal using
add_SNo_rotate_4_1 (v * z ) (x * w ) (u * w ) (y * w ) Lvz Lxw Luw Lyw (from left to right).
We will
prove y * w + v * z + x * w + u * w ≤ x * z + u * z + v * w + y * w .
rewrite the current goal using
add_SNo_rotate_4_1 (x * z ) (u * z ) (v * w ) (y * w ) Lxz Luz Lvw Lyw (from left to right).
We will
prove y * w + v * z + x * w + u * w ≤ y * w + x * z + u * z + v * w .
We will
prove v * z + x * w + u * w ≤ x * z + u * z + v * w .
rewrite the current goal using
IHyz u (SNoR_SNoS y Hy u Hu ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + (x + u ) * w ≤ x * z + u * z + v * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (u * z ) (v * w ) Lxz Luz Lvw (from left to right).
We will
prove v * z + (x + u ) * w ≤ (x * z + u * z ) + v * w .
rewrite the current goal using
IHy u (SNoR_SNoS y Hy u Hu ) (from right to left).
We will
prove v * z + (x + u ) * w ≤ (x + u ) * z + v * w .
Apply mul_SNo_Le v w (x + u ) z Hv1 Hw1 Lxu Hz to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
We will
prove x * z + u * z + v * w + y * w < x * z + y * z + v * w + u * w .
We will
prove u * z + v * w + y * w < y * z + v * w + u * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (y * w ) Luz Lvw Lyw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (y * z ) (v * w ) (u * w ) Lyz Lvw Luw (from left to right).
We will
prove v * w + u * z + y * w < v * w + y * z + u * w .
We will
prove u * z + y * w < y * z + u * w .
rewrite the current goal using
add_SNo_com (y * z ) (u * w ) Lyz Luw (from left to right).
rewrite the current goal using
add_SNo_com (u * z ) (y * w ) Luz Lyw (from left to right).
We will
prove y * w + u * z < u * w + y * z .
An
exact proof term for the current goal is
mul_SNo_Lt u w y z Hu1 Hw1 Hy Hz Hu3 Hw3 .
Let u be given.
rewrite the current goal using LE (from right to left).
We will
prove x * z + y * z < u .
Apply HRE u Hu to the current goal.
Let v be given.
Let w be given.
rewrite the current goal using Hue (from left to right).
We will
prove x * z + y * z < v * z + (x + y ) * w + - v * w .
Apply SNoL_E (x + y ) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lxwyw :
SNo (x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * w ) (y * w ) Lxw Lyw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvzxwyw :
SNo (v * z + x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (v * z ) (x * w ) (y * w ) Lvz Lxw Lyw .
We prove the intermediate
claim Lxzyzvw :
SNo (x * z + y * z + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw .
We will
prove x * z + y * z < v * z + (x + y ) * w + - v * w .
We will
prove (x * z + y * z ) + v * w < v * z + (x + y ) * w .
rewrite the current goal using
IHz w (SNoR_SNoS z Hz w Hw ) (from left to right).
We will
prove (x * z + y * z ) + v * w < v * z + x * w + y * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw (from right to left).
We will
prove x * z + y * z + v * w < v * z + x * w + y * w .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoL_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Luy :
SNo (u + y ) .
An
exact proof term for the current goal is
SNo_add_SNo u y Hu1 Hy .
We will
prove (x * z + y * z + v * w ) + u * w < (v * z + x * w + y * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w .
We will
prove x * z + y * z + v * w + u * w < u * z + y * z + v * w + x * w .
We will
prove y * z + x * z + v * w + u * w < y * z + u * z + v * w + x * w .
We will
prove x * z + v * w + u * w < u * z + v * w + x * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (x * w ) Luz Lvw Lxw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (x * z ) (v * w ) (u * w ) Lxz Lvw Luw (from left to right).
We will
prove v * w + x * z + u * w < v * w + u * z + x * w .
We will
prove x * z + u * w < u * z + x * w .
rewrite the current goal using
add_SNo_com (x * z ) (u * w ) Lxz Luw (from left to right).
rewrite the current goal using
add_SNo_com (u * z ) (x * w ) Luz Lxw (from left to right).
An
exact proof term for the current goal is
mul_SNo_Lt x w u z Hx Hw1 Hu1 Hz Hu3 Hw3 .
We will
prove u * z + y * z + v * w + x * w ≤ v * z + x * w + y * w + u * w .
We will
prove u * z + y * z + v * w + x * w ≤ x * w + v * z + y * w + u * w .
rewrite the current goal using
add_SNo_rotate_4_1 (u * z ) (y * z ) (v * w ) (x * w ) Luz Lyz Lvw Lxw (from left to right).
We will
prove x * w + u * z + y * z + v * w ≤ x * w + v * z + y * w + u * w .
We will
prove u * z + y * z + v * w ≤ v * z + y * w + u * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove u * z + y * z + v * w ≤ v * z + u * w + y * w .
rewrite the current goal using
IHxz u (SNoL_SNoS x Hx u Hu ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove u * z + y * z + v * w ≤ v * z + (u + y ) * w .
rewrite the current goal using
add_SNo_assoc (u * z ) (y * z ) (v * w ) Luz Lyz Lvw (from left to right).
We will
prove (u * z + y * z ) + v * w ≤ v * z + (u + y ) * w .
rewrite the current goal using
IHx u (SNoL_SNoS x Hx u Hu ) (from right to left).
We will
prove (u + y ) * z + v * w ≤ v * z + (u + y ) * w .
We will
prove v * w + (u + y ) * z ≤ (u + y ) * w + v * z .
Apply mul_SNo_Le (u + y ) w v z Luy Hw1 Hv1 Hz to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoL_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Lxu :
SNo (x + u ) .
An
exact proof term for the current goal is
SNo_add_SNo x u Hx Hu1 .
We will
prove (x * z + y * z + v * w ) + u * w < (v * z + x * w + y * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w .
We will
prove x * z + y * z + v * w + u * w < x * z + u * z + v * w + y * w .
We will
prove y * z + v * w + u * w < u * z + v * w + y * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (y * w ) Luz Lvw Lyw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (y * z ) (v * w ) (u * w ) Lyz Lvw Luw (from left to right).
We will
prove v * w + y * z + u * w < v * w + u * z + y * w .
We will
prove y * z + u * w < u * z + y * w .
rewrite the current goal using
add_SNo_com (y * z ) (u * w ) Lyz Luw (from left to right).
rewrite the current goal using
add_SNo_com (u * z ) (y * w ) Luz Lyw (from left to right).
An
exact proof term for the current goal is
mul_SNo_Lt y w u z Hy Hw1 Hu1 Hz Hu3 Hw3 .
We will
prove x * z + u * z + v * w + y * w ≤ v * z + x * w + y * w + u * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove x * z + u * z + v * w + y * w ≤ v * z + x * w + u * w + y * w .
rewrite the current goal using
add_SNo_rotate_4_1 (v * z ) (x * w ) (u * w ) (y * w ) Lvz Lxw Luw Lyw (from left to right).
We will
prove x * z + u * z + v * w + y * w ≤ y * w + v * z + x * w + u * w .
rewrite the current goal using
add_SNo_rotate_4_1 (x * z ) (u * z ) (v * w ) (y * w ) Lxz Luz Lvw Lyw (from left to right).
We will
prove y * w + x * z + u * z + v * w ≤ y * w + v * z + x * w + u * w .
We will
prove x * z + u * z + v * w ≤ v * z + x * w + u * w .
rewrite the current goal using
IHyz u (SNoL_SNoS y Hy u Hu ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove x * z + u * z + v * w ≤ v * z + (x + u ) * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (u * z ) (v * w ) Lxz Luz Lvw (from left to right).
We will
prove (x * z + u * z ) + v * w ≤ v * z + (x + u ) * w .
rewrite the current goal using
IHy u (SNoL_SNoS y Hy u Hu ) (from right to left).
We will
prove (x + u ) * z + v * w ≤ v * z + (x + u ) * w .
We will
prove v * w + (x + u ) * z ≤ (x + u ) * w + v * z .
Apply mul_SNo_Le (x + u ) w v z Lxu Hw1 Hv1 Hz to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
Let v be given.
Let w be given.
rewrite the current goal using Hue (from left to right).
We will
prove x * z + y * z < v * z + (x + y ) * w + - v * w .
Apply SNoR_E (x + y ) Lxy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lxwyw :
SNo (x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * w ) (y * w ) Lxw Lyw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvzxwyw :
SNo (v * z + x * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (v * z ) (x * w ) (y * w ) Lvz Lxw Lyw .
We prove the intermediate
claim Lxzyzvw :
SNo (x * z + y * z + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw .
We will
prove x * z + y * z < v * z + (x + y ) * w + - v * w .
We will
prove (x * z + y * z ) + v * w < v * z + (x + y ) * w .
rewrite the current goal using
IHz w (SNoL_SNoS z Hz w Hw ) (from left to right).
We will
prove (x * z + y * z ) + v * w < v * z + x * w + y * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (y * z ) (v * w ) Lxz Lyz Lvw (from right to left).
We will
prove x * z + y * z + v * w < v * z + x * w + y * w .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoR_E x Hx u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Luy :
SNo (u + y ) .
An
exact proof term for the current goal is
SNo_add_SNo u y Hu1 Hy .
We will
prove (x * z + y * z + v * w ) + u * w < (v * z + x * w + y * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w .
We will
prove x * z + y * z + v * w + u * w < u * z + y * z + v * w + x * w .
We will
prove y * z + x * z + v * w + u * w < y * z + u * z + v * w + x * w .
We will
prove x * z + v * w + u * w < u * z + v * w + x * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (x * w ) Luz Lvw Lxw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (x * z ) (v * w ) (u * w ) Lxz Lvw Luw (from left to right).
We will
prove v * w + x * z + u * w < v * w + u * z + x * w .
We will
prove x * z + u * w < u * z + x * w .
An
exact proof term for the current goal is
mul_SNo_Lt u z x w Hu1 Hz Hx Hw1 Hu3 Hw3 .
We will
prove u * z + y * z + v * w + x * w ≤ v * z + x * w + y * w + u * w .
We will
prove u * z + y * z + v * w + x * w ≤ x * w + v * z + y * w + u * w .
rewrite the current goal using
add_SNo_rotate_4_1 (u * z ) (y * z ) (v * w ) (x * w ) Luz Lyz Lvw Lxw (from left to right).
We will
prove x * w + u * z + y * z + v * w ≤ x * w + v * z + y * w + u * w .
We will
prove u * z + y * z + v * w ≤ v * z + y * w + u * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove u * z + y * z + v * w ≤ v * z + u * w + y * w .
rewrite the current goal using
IHxz u (SNoR_SNoS x Hx u Hu ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove u * z + y * z + v * w ≤ v * z + (u + y ) * w .
rewrite the current goal using
add_SNo_assoc (u * z ) (y * z ) (v * w ) Luz Lyz Lvw (from left to right).
We will
prove (u * z + y * z ) + v * w ≤ v * z + (u + y ) * w .
rewrite the current goal using
IHx u (SNoR_SNoS x Hx u Hu ) (from right to left).
We will
prove (u + y ) * z + v * w ≤ v * z + (u + y ) * w .
Apply mul_SNo_Le v z (u + y ) w Hv1 Hz Luy Hw1 to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
Assume H1 .
Apply H1 to the current goal.
Let u be given.
Assume H1 .
Apply H1 to the current goal.
Apply SNoR_E y Hy u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
We prove the intermediate
claim Luw :
SNo (u * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo u w Hu1 Hw1 .
We prove the intermediate
claim Luz :
SNo (u * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo u z Hu1 Hz .
We prove the intermediate
claim Lxu :
SNo (x + u ) .
An
exact proof term for the current goal is
SNo_add_SNo x u Hx Hu1 .
We will
prove (x * z + y * z + v * w ) + u * w < (v * z + x * w + y * w ) + u * w .
rewrite the current goal using
add_SNo_assoc_4 (v * z ) (x * w ) (y * w ) (u * w ) Lvz Lxw Lyw Luw (from right to left).
rewrite the current goal using
add_SNo_assoc_4 (x * z ) (y * z ) (v * w ) (u * w ) Lxz Lyz Lvw Luw (from right to left).
We will
prove x * z + y * z + v * w + u * w < v * z + x * w + y * w + u * w .
We will
prove x * z + y * z + v * w + u * w < x * z + u * z + v * w + y * w .
We will
prove y * z + v * w + u * w < u * z + v * w + y * w .
rewrite the current goal using
add_SNo_com_3_0_1 (u * z ) (v * w ) (y * w ) Luz Lvw Lyw (from left to right).
rewrite the current goal using
add_SNo_com_3_0_1 (y * z ) (v * w ) (u * w ) Lyz Lvw Luw (from left to right).
We will
prove v * w + y * z + u * w < v * w + u * z + y * w .
We will
prove y * z + u * w < u * z + y * w .
An
exact proof term for the current goal is
mul_SNo_Lt u z y w Hu1 Hz Hy Hw1 Hu3 Hw3 .
We will
prove x * z + u * z + v * w + y * w ≤ v * z + x * w + y * w + u * w .
rewrite the current goal using
add_SNo_com (y * w ) (u * w ) Lyw Luw (from left to right).
We will
prove x * z + u * z + v * w + y * w ≤ v * z + x * w + u * w + y * w .
rewrite the current goal using
add_SNo_rotate_4_1 (v * z ) (x * w ) (u * w ) (y * w ) Lvz Lxw Luw Lyw (from left to right).
We will
prove x * z + u * z + v * w + y * w ≤ y * w + v * z + x * w + u * w .
rewrite the current goal using
add_SNo_rotate_4_1 (x * z ) (u * z ) (v * w ) (y * w ) Lxz Luz Lvw Lyw (from left to right).
We will
prove y * w + x * z + u * z + v * w ≤ y * w + v * z + x * w + u * w .
We will
prove x * z + u * z + v * w ≤ v * z + x * w + u * w .
rewrite the current goal using
IHyz u (SNoR_SNoS y Hy u Hu ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove x * z + u * z + v * w ≤ v * z + (x + u ) * w .
rewrite the current goal using
add_SNo_assoc (x * z ) (u * z ) (v * w ) Lxz Luz Lvw (from left to right).
We will
prove (x * z + u * z ) + v * w ≤ v * z + (x + u ) * w .
rewrite the current goal using
IHy u (SNoR_SNoS y Hy u Hu ) (from right to left).
We will
prove (x + u ) * z + v * w ≤ v * z + (x + u ) * w .
Apply mul_SNo_Le v z (x + u ) w Hv1 Hz Lxu Hw1 to the current goal.
An exact proof term for the current goal is Hvu .
An exact proof term for the current goal is Hw3 .
Let u' be given.
rewrite the current goal using HE (from right to left).
We will
prove u' < (x + y ) * z .
Apply binunionE L1 L2 u' Hu' to the current goal.
Let u be given.
rewrite the current goal using Hu'u (from left to right).
We will
prove u + y * z < (x + y ) * z .
Apply SNoL_E (x * z ) Lxz u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
Let v be given.
Let w be given.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvy :
SNo (v + y ) .
An
exact proof term for the current goal is
SNo_add_SNo v y Hv1 Hy .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luyz :
SNo (u + y * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (y * z ) Hu1 Lyz .
We prove the intermediate
claim Lvwyw :
SNo (v * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (y * w ) Lvw Lyw .
We prove the intermediate
claim Lvzxw :
SNo (v * z + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * z ) (x * w ) Lvz Lxw .
We will
prove u + y * z < (x + y ) * z .
We will
prove (u + y * z ) + v * w + y * w < (x + y ) * z + v * w + y * w .
We will
prove (u + v * w ) + y * z + y * w < (x + y ) * z + v * w + y * w .
We will
prove (u + v * w ) + y * z + y * w ≤ v * z + y * z + x * w + y * w .
We will
prove (u + v * w ) + y * z + y * w ≤ (v * z + y * z ) + x * w + y * w .
We will
prove (u + v * w ) + y * z + y * w ≤ (v * z + x * w ) + y * z + y * w .
We will
prove u + v * w ≤ v * z + x * w .
An exact proof term for the current goal is Hvw .
We will
prove v * z + y * z + x * w + y * w < (x + y ) * z + v * w + y * w .
rewrite the current goal using
IHz w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + y * z + (x + y ) * w < (x + y ) * z + v * w + y * w .
rewrite the current goal using
IHxz v (SNoL_SNoS x Hx v Hv ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + y * z + (x + y ) * w < (x + y ) * z + (v + y ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (y * z ) ((x + y ) * w ) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using
IHx v (SNoL_SNoS x Hx v Hv ) (from right to left).
We will
prove (v + y ) * z + (x + y ) * w < (x + y ) * z + (v + y ) * w .
Apply mul_SNo_Lt (x + y ) z (v + y ) w Lxy Hz Lvy Hw1 to the current goal.
We will
prove v + y < x + y .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
Let v be given.
Let w be given.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvy :
SNo (v + y ) .
An
exact proof term for the current goal is
SNo_add_SNo v y Hv1 Hy .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luyz :
SNo (u + y * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (y * z ) Hu1 Lyz .
We prove the intermediate
claim Lvwyw :
SNo (v * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (y * w ) Lvw Lyw .
We will
prove u + y * z < (x + y ) * z .
We will
prove (u + y * z ) + v * w + y * w < (x + y ) * z + v * w + y * w .
We will
prove (u + v * w ) + y * z + y * w < (x + y ) * z + v * w + y * w .
We will
prove (u + v * w ) + y * z + y * w ≤ v * z + y * z + x * w + y * w .
We will
prove (u + v * w ) + y * z + y * w ≤ (v * z + y * z ) + x * w + y * w .
We will
prove (u + v * w ) + y * z + y * w ≤ (v * z + x * w ) + y * z + y * w .
We will
prove u + v * w ≤ v * z + x * w .
An exact proof term for the current goal is Hvw .
We will
prove v * z + y * z + x * w + y * w < (x + y ) * z + v * w + y * w .
rewrite the current goal using
IHz w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + y * z + (x + y ) * w < (x + y ) * z + v * w + y * w .
rewrite the current goal using
IHxz v (SNoR_SNoS x Hx v Hv ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + y * z + (x + y ) * w < (x + y ) * z + (v + y ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (y * z ) ((x + y ) * w ) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using
IHx v (SNoR_SNoS x Hx v Hv ) (from right to left).
We will
prove (v + y ) * z + (x + y ) * w < (x + y ) * z + (v + y ) * w .
Apply mul_SNo_Lt (v + y ) w (x + y ) z Lvy Hw1 Lxy Hz to the current goal.
We will
prove x + y < v + y .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
Let u be given.
rewrite the current goal using Hu'u (from left to right).
We will
prove x * z + u < (x + y ) * z .
Apply SNoL_E (y * z ) Lyz u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
rewrite the current goal using
add_SNo_com (x * z ) u Lxz Hu1 (from left to right).
We will
prove u + x * z < (x + y ) * z .
Let v be given.
Let w be given.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lxv :
SNo (x + v ) .
An
exact proof term for the current goal is
SNo_add_SNo x v Hx Hv1 .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luxz :
SNo (u + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (x * z ) Hu1 Lxz .
We prove the intermediate
claim Lvwxw :
SNo (v * w + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (x * w ) Lvw Lxw .
We will
prove u + x * z < (x + y ) * z .
We will
prove (u + x * z ) + v * w + x * w < (x + y ) * z + v * w + x * w .
We will
prove (u + v * w ) + x * z + x * w < (x + y ) * z + v * w + x * w .
We will
prove (u + v * w ) + x * z + x * w ≤ v * z + x * z + x * w + y * w .
We will
prove (u + v * w ) + x * z + x * w ≤ (v * z + x * z ) + x * w + y * w .
rewrite the current goal using
add_SNo_com (x * w ) (y * w ) Lxw Lyw (from left to right).
We will
prove (u + v * w ) + x * z + x * w ≤ (v * z + x * z ) + y * w + x * w .
We will
prove (u + v * w ) + x * z + x * w ≤ (v * z + y * w ) + x * z + x * w .
We will
prove u + v * w ≤ v * z + y * w .
An exact proof term for the current goal is Hvw .
We will
prove v * z + x * z + x * w + y * w < (x + y ) * z + v * w + x * w .
rewrite the current goal using
IHz w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + x * z + (x + y ) * w < (x + y ) * z + v * w + x * w .
rewrite the current goal using
add_SNo_com (v * w ) (x * w ) Lvw Lxw (from left to right).
rewrite the current goal using
IHyz v (SNoL_SNoS y Hy v Hv ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + x * z + (x + y ) * w < (x + y ) * z + (x + v ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (x * z ) ((x + y ) * w ) Lvz Lxz Lxyw (from left to right).
We will
prove (v * z + x * z ) + (x + y ) * w < (x + y ) * z + (x + v ) * w .
rewrite the current goal using
add_SNo_com (v * z ) (x * z ) Lvz Lxz (from left to right).
We will
prove (x * z + v * z ) + (x + y ) * w < (x + y ) * z + (x + v ) * w .
rewrite the current goal using
IHy v (SNoL_SNoS y Hy v Hv ) (from right to left).
We will
prove (x + v ) * z + (x + y ) * w < (x + y ) * z + (x + v ) * w .
Apply mul_SNo_Lt (x + y ) z (x + v ) w Lxy Hz Lxv Hw1 to the current goal.
We will
prove x + v < x + y .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
Let v be given.
Let w be given.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lxv :
SNo (x + v ) .
An
exact proof term for the current goal is
SNo_add_SNo x v Hx Hv1 .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luxz :
SNo (u + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (x * z ) Hu1 Lxz .
We prove the intermediate
claim Lvwxw :
SNo (v * w + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (x * w ) Lvw Lxw .
We prove the intermediate
claim Lvzxw :
SNo (v * z + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * z ) (x * w ) Lvz Lxw .
We will
prove u + x * z < (x + y ) * z .
We will
prove (u + x * z ) + v * w + x * w < (x + y ) * z + v * w + x * w .
We will
prove (u + v * w ) + x * z + x * w < (x + y ) * z + v * w + x * w .
We will
prove (u + v * w ) + x * z + x * w ≤ v * z + x * z + x * w + y * w .
rewrite the current goal using
add_SNo_com (y * w ) (x * w ) Lyw Lxw (from right to left).
We will
prove (u + v * w ) + x * z + x * w ≤ (v * z + x * z ) + y * w + x * w .
We will
prove (u + v * w ) + x * z + x * w ≤ (v * z + y * w ) + x * z + x * w .
We will
prove u + v * w ≤ v * z + y * w .
An exact proof term for the current goal is Hvw .
We will
prove v * z + x * z + x * w + y * w < (x + y ) * z + v * w + x * w .
rewrite the current goal using
IHz w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + x * z + (x + y ) * w < (x + y ) * z + v * w + x * w .
rewrite the current goal using
add_SNo_com (v * w ) (x * w ) Lvw Lxw (from left to right).
We will
prove v * z + x * z + (x + y ) * w < (x + y ) * z + x * w + v * w .
rewrite the current goal using
IHyz v (SNoR_SNoS y Hy v Hv ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove v * z + x * z + (x + y ) * w < (x + y ) * z + (x + v ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (x * z ) ((x + y ) * w ) Lvz Lxz Lxyw (from left to right).
We will
prove (v * z + x * z ) + (x + y ) * w < (x + y ) * z + (x + v ) * w .
rewrite the current goal using
add_SNo_com (v * z ) (x * z ) Lvz Lxz (from left to right).
rewrite the current goal using
IHy v (SNoR_SNoS y Hy v Hv ) (from right to left).
We will
prove (x + v ) * z + (x + y ) * w < (x + y ) * z + (x + v ) * w .
We will
prove (x + y ) * w + (x + v ) * z < (x + v ) * w + (x + y ) * z .
Apply mul_SNo_Lt (x + v ) w (x + y ) z Lxv Hw1 Lxy Hz to the current goal.
We will
prove x + y < x + v .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
Let u' be given.
rewrite the current goal using HE (from right to left).
We will
prove (x + y ) * z < u' .
Apply binunionE R1 R2 u' Hu' to the current goal.
Let u be given.
rewrite the current goal using Hu'u (from left to right).
We will
prove (x + y ) * z < u + y * z .
Apply SNoR_E (x * z ) Lxz u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
Let v be given.
Let w be given.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvy :
SNo (v + y ) .
An
exact proof term for the current goal is
SNo_add_SNo v y Hv1 Hy .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luyz :
SNo (u + y * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (y * z ) Hu1 Lyz .
We prove the intermediate
claim Lvwyw :
SNo (v * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (y * w ) Lvw Lyw .
We prove the intermediate
claim Lvzxw :
SNo (v * z + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * z ) (x * w ) Lvz Lxw .
We will
prove (x + y ) * z < u + y * z .
We will
prove (x + y ) * z + v * w + y * w < (u + y * z ) + v * w + y * w .
We will
prove (x + y ) * z + v * w + y * w < (u + v * w ) + y * z + y * w .
We will
prove (x + y ) * z + v * w + y * w < v * z + y * z + x * w + y * w .
rewrite the current goal using
IHz w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + v * w + y * w < v * z + y * z + (x + y ) * w .
rewrite the current goal using
IHxz v (SNoL_SNoS x Hx v Hv ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + (v + y ) * w < v * z + y * z + (x + y ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (y * z ) ((x + y ) * w ) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using
IHx v (SNoL_SNoS x Hx v Hv ) (from right to left).
We will
prove (x + y ) * z + (v + y ) * w < (v + y ) * z + (x + y ) * w .
We will
prove (v + y ) * w + (x + y ) * z < (x + y ) * w + (v + y ) * z .
Apply mul_SNo_Lt (x + y ) w (v + y ) z Lxy Hw1 Lvy Hz to the current goal.
We will
prove v + y < x + y .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
We will
prove v * z + y * z + x * w + y * w ≤ (u + v * w ) + y * z + y * w .
We will
prove (v * z + y * z ) + x * w + y * w ≤ (u + v * w ) + y * z + y * w .
We will
prove (v * z + x * w ) + y * z + y * w ≤ (u + v * w ) + y * z + y * w .
We will
prove v * z + x * w ≤ u + v * w .
An exact proof term for the current goal is Hvw .
Let v be given.
Let w be given.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lvy :
SNo (v + y ) .
An
exact proof term for the current goal is
SNo_add_SNo v y Hv1 Hy .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luyz :
SNo (u + y * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (y * z ) Hu1 Lyz .
We prove the intermediate
claim Lvwyw :
SNo (v * w + y * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (y * w ) Lvw Lyw .
We will
prove (x + y ) * z < u + y * z .
We will
prove (x + y ) * z + v * w + y * w < (u + y * z ) + v * w + y * w .
We will
prove (x + y ) * z + v * w + y * w < (u + v * w ) + y * z + y * w .
We will
prove (x + y ) * z + v * w + y * w < v * z + y * z + x * w + y * w .
rewrite the current goal using
IHz w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + v * w + y * w < v * z + y * z + (x + y ) * w .
rewrite the current goal using
IHxz v (SNoR_SNoS x Hx v Hv ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + (v + y ) * w < v * z + y * z + (x + y ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (y * z ) ((x + y ) * w ) Lvz Lyz Lxyw (from left to right).
rewrite the current goal using
IHx v (SNoR_SNoS x Hx v Hv ) (from right to left).
We will
prove (x + y ) * z + (v + y ) * w < (v + y ) * z + (x + y ) * w .
Apply mul_SNo_Lt (v + y ) z (x + y ) w Lvy Hz Lxy Hw1 to the current goal.
We will
prove x + y < v + y .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
We will
prove v * z + y * z + x * w + y * w ≤ (u + v * w ) + y * z + y * w .
We will
prove (v * z + y * z ) + x * w + y * w ≤ (u + v * w ) + y * z + y * w .
We will
prove (v * z + x * w ) + y * z + y * w ≤ (u + v * w ) + y * z + y * w .
We will
prove v * z + x * w ≤ u + v * w .
An exact proof term for the current goal is Hvw .
Let u be given.
rewrite the current goal using Hu'u (from left to right).
We will
prove (x + y ) * z < x * z + u .
Apply SNoR_E (y * z ) Lyz u Hu to the current goal.
Assume Hu1 Hu2 Hu3 .
rewrite the current goal using
add_SNo_com (x * z ) u Lxz Hu1 (from left to right).
We will
prove (x + y ) * z < u + x * z .
Let v be given.
Let w be given.
Apply SNoL_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoR_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lxv :
SNo (x + v ) .
An
exact proof term for the current goal is
SNo_add_SNo x v Hx Hv1 .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luxz :
SNo (u + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (x * z ) Hu1 Lxz .
We prove the intermediate
claim Lvwxw :
SNo (v * w + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (x * w ) Lvw Lxw .
We will
prove (x + y ) * z < u + x * z .
We will
prove (x + y ) * z + v * w + x * w < (u + x * z ) + v * w + x * w .
We will
prove (x + y ) * z + v * w + x * w < (u + v * w ) + x * z + x * w .
We will
prove (x + y ) * z + v * w + x * w < v * z + x * z + x * w + y * w .
rewrite the current goal using
IHz w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + v * w + x * w < v * z + x * z + (x + y ) * w .
rewrite the current goal using
add_SNo_com (v * w ) (x * w ) Lvw Lxw (from left to right).
rewrite the current goal using
IHyz v (SNoL_SNoS y Hy v Hv ) w (SNoR_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + (x + v ) * w < v * z + x * z + (x + y ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (x * z ) ((x + y ) * w ) Lvz Lxz Lxyw (from left to right).
We will
prove (x + y ) * z + (x + v ) * w < (v * z + x * z ) + (x + y ) * w .
rewrite the current goal using
add_SNo_com (v * z ) (x * z ) Lvz Lxz (from left to right).
We will
prove (x + y ) * z + (x + v ) * w < (x * z + v * z ) + (x + y ) * w .
rewrite the current goal using
IHy v (SNoL_SNoS y Hy v Hv ) (from right to left).
We will
prove (x + y ) * z + (x + v ) * w < (x + v ) * z + (x + y ) * w .
We will
prove (x + v ) * w + (x + y ) * z < (x + y ) * w + (x + v ) * z .
Apply mul_SNo_Lt (x + y ) w (x + v ) z Lxy Hw1 Lxv Hz to the current goal.
We will
prove x + v < x + y .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
We will
prove v * z + x * z + x * w + y * w ≤ (u + v * w ) + x * z + x * w .
We will
prove (v * z + x * z ) + x * w + y * w ≤ (u + v * w ) + x * z + x * w .
rewrite the current goal using
add_SNo_com (x * w ) (y * w ) Lxw Lyw (from left to right).
We will
prove (v * z + x * z ) + y * w + x * w ≤ (u + v * w ) + x * z + x * w .
We will
prove (v * z + y * w ) + x * z + x * w ≤ (u + v * w ) + x * z + x * w .
We will
prove v * z + y * w ≤ u + v * w .
An exact proof term for the current goal is Hvw .
Let v be given.
Let w be given.
Apply SNoR_E y Hy v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
Apply SNoL_E z Hz w Hw to the current goal.
Assume Hw1 Hw2 Hw3 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Hv1 Hw1 .
We prove the intermediate
claim Lxv :
SNo (x + v ) .
An
exact proof term for the current goal is
SNo_add_SNo x v Hx Hv1 .
We prove the intermediate
claim Luvw :
SNo (u + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo u (v * w ) Hu1 Lvw .
We prove the intermediate
claim Lxyw :
SNo ((x + y ) * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x + y ) w Lxy Hw1 .
We prove the intermediate
claim Lvz :
SNo (v * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo v z Hv1 Hz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo x w Hx Hw1 .
We prove the intermediate
claim Lyw :
SNo (y * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo y w Hy Hw1 .
We prove the intermediate
claim Luxz :
SNo (u + x * z ) .
An
exact proof term for the current goal is
SNo_add_SNo u (x * z ) Hu1 Lxz .
We prove the intermediate
claim Lvwxw :
SNo (v * w + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (x * w ) Lvw Lxw .
We prove the intermediate
claim Lvzxw :
SNo (v * z + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * z ) (x * w ) Lvz Lxw .
We will
prove (x + y ) * z < u + x * z .
We will
prove (x + y ) * z + v * w + x * w < (u + x * z ) + v * w + x * w .
We will
prove (x + y ) * z + v * w + x * w < (u + v * w ) + x * z + x * w .
We will
prove (x + y ) * z + v * w + x * w < v * z + x * z + x * w + y * w .
rewrite the current goal using
IHz w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + v * w + x * w < v * z + x * z + (x + y ) * w .
rewrite the current goal using
add_SNo_com (v * w ) (x * w ) Lvw Lxw (from left to right).
We will
prove (x + y ) * z + x * w + v * w < v * z + x * z + (x + y ) * w .
rewrite the current goal using
IHyz v (SNoR_SNoS y Hy v Hv ) w (SNoL_SNoS z Hz w Hw ) (from right to left).
We will
prove (x + y ) * z + (x + v ) * w < v * z + x * z + (x + y ) * w .
rewrite the current goal using
add_SNo_assoc (v * z ) (x * z ) ((x + y ) * w ) Lvz Lxz Lxyw (from left to right).
We will
prove (x + y ) * z + (x + v ) * w < (v * z + x * z ) + (x + y ) * w .
rewrite the current goal using
add_SNo_com (v * z ) (x * z ) Lvz Lxz (from left to right).
rewrite the current goal using
IHy v (SNoR_SNoS y Hy v Hv ) (from right to left).
We will
prove (x + y ) * z + (x + v ) * w < (x + v ) * z + (x + y ) * w .
We will
prove (x + v ) * w + (x + y ) * z < (x + y ) * w + (x + v ) * z .
We will
prove (x + y ) * z + (x + v ) * w < (x + v ) * z + (x + y ) * w .
Apply mul_SNo_Lt (x + v ) z (x + y ) w Lxv Hz Lxy Hw1 to the current goal.
We will
prove x + y < x + v .
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw3 .
We will
prove v * z + x * z + x * w + y * w ≤ (u + v * w ) + x * z + x * w .
rewrite the current goal using
add_SNo_com (y * w ) (x * w ) Lyw Lxw (from right to left).
We will
prove (v * z + x * z ) + y * w + x * w ≤ (u + v * w ) + x * z + x * w .
We will
prove (v * z + y * w ) + x * z + x * w ≤ (u + v * w ) + x * z + x * w .
We will
prove v * z + y * w ≤ u + v * w .
An exact proof term for the current goal is Hvw .
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz .
Use transitivity with
((y + z ) * x ) , and
(y * x + z * x ) .
An
exact proof term for the current goal is
mul_SNo_distrR y z x Hy Hz Hx .
We will
prove y * x + z * x = x * y + x * z .
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_com y x Hy Hx .
An
exact proof term for the current goal is
mul_SNo_com z x Hz Hx .
∎
Beginning of Section mul_SNo_assoc_lems
Variable M : set → set → set
Notation . We use
* as an infix operator with priority 355 and which associates to the right corresponding to applying term
M .
Hypothesis SNo_M : ∀x y, SNo x → SNo y → SNo (x * y )
Hypothesis DL : ∀x y z, SNo x → SNo y → SNo z → x * (y + z ) = x * y + x * z
Hypothesis DR : ∀x y z, SNo x → SNo y → SNo z → (x + y ) * z = x * z + y * z
Hypothesis M_Lt : ∀x y u v, SNo x → SNo y → SNo u → SNo v → u < x → v < y → u * y + x * v < x * y + u * v
Hypothesis M_Le : ∀x y u v, SNo x → SNo y → SNo u → SNo v → u ≤ x → v ≤ y → u * y + x * v ≤ x * y + u * v
Theorem. (
mul_SNo_assoc_lem1 ) The following is provable:
∀x y z, SNo x → SNo y → SNo z → (∀ u ∈ SNoS_ (SNoLev x ) , u * (y * z ) = (u * y ) * z ) → (∀ v ∈ SNoS_ (SNoLev y ) , x * (v * z ) = (x * v ) * z ) → (∀ w ∈ SNoS_ (SNoLev z ) , x * (y * w ) = (x * y ) * w ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ v ∈ SNoS_ (SNoLev y ) , u * (v * z ) = (u * v ) * z ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ w ∈ SNoS_ (SNoLev z ) , u * (y * w ) = (u * y ) * w ) → (∀ v ∈ SNoS_ (SNoLev y ) , ∀ w ∈ SNoS_ (SNoLev z ) , x * (v * w ) = (x * v ) * w ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ v ∈ SNoS_ (SNoLev y ) , ∀ w ∈ SNoS_ (SNoLev z ) , u * (v * w ) = (u * v ) * w ) → ∀L, (∀ u ∈ L , ∀q : prop , (∀ v ∈ SNoL x , ∀ w ∈ SNoL (y * z ) , u = v * (y * z ) + x * w + - v * w → q ) → (∀ v ∈ SNoR x , ∀ w ∈ SNoR (y * z ) , u = v * (y * z ) + x * w + - v * w → q ) → q ) → ∀ u ∈ L , u < (x * y ) * z
Proof: Let x, y and z be given.
Assume Hx Hy Hz .
Assume IHx IHy IHz IHxy IHxz IHyz IHxyz .
Let L be given.
Assume HLE .
Let u be given.
We will
prove u < (x * y ) * z .
We prove the intermediate
claim Lxy :
SNo (x * y ) .
An exact proof term for the current goal is SNo_M x y Hx Hy .
We prove the intermediate
claim Lyz :
SNo (y * z ) .
An exact proof term for the current goal is SNo_M y z Hy Hz .
We prove the intermediate
claim Lxyz2 :
SNo ((x * y ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y ) z Lxy Hz .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
Let w1 be given.
Assume Hw1 .
Let w2 be given.
Assume Hw2 .
Assume Hue H1 H2 .
Assume Hv1 Hv2 Hv3 Hv4 .
Assume Hw11 Hw12 Hw13 Hw14 .
Assume Hw21 Hw22 Hw23 Hw24 .
We prove the intermediate
claim Lvyz :
SNo (v * (y * z ) ) .
An
exact proof term for the current goal is
SNo_M v (y * z ) Hv3 Lyz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An exact proof term for the current goal is SNo_M x w Hx Hw .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw13 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An exact proof term for the current goal is SNo_M v w Hv3 Hw .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv3 Hw13 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv3 Hy .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw13 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw23 .
We prove the intermediate
claim Lvw1z :
SNo (v * (w1 * z ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z ) Hv3 Lw1z .
We prove the intermediate
claim Lvyw2 :
SNo (v * (y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (y * w2 ) Hv3 Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw13 Hw23 .
We prove the intermediate
claim Lxw1w2 :
SNo (x * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * w2 ) Hx Lw1w2 .
We prove the intermediate
claim Lxw1z :
SNo (x * (w1 * z ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z ) Hx Lw1z .
We prove the intermediate
claim Lxyw2 :
SNo (x * (y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (y * w2 ) Hx Lyw2 .
We prove the intermediate
claim Lvyzxw :
SNo (v * (y * z ) + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * (y * z ) ) (x * w ) Lvyz Lxw .
We prove the intermediate
claim Lxyzvw :
SNo ((x * y ) * z + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo ((x * y ) * z ) (v * w ) Lxyz2 Lvw .
We prove the intermediate
claim Lvw1w2 :
SNo (v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * w2 ) Hv3 Lw1w2 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Hw Lw1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv3 Lww1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv3 (SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 ) .
We prove the intermediate
claim Lvwvw1w2 :
SNo (v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (v * (w1 * w2 ) ) Lvw Lvw1w2 .
We prove the intermediate
claim Lvyzxw1zxyw2vwvw1w2 :
SNo (v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo_4 (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) ) (v * w + v * (w1 * w2 ) ) Lvyz Lxw1z Lxyw2 Lvwvw1w2 .
We prove the intermediate
claim Lvw1zvyw2xw1w2 :
SNo (v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (v * (w1 * z ) ) (v * (y * w2 ) ) (x * (w1 * w2 ) ) Lvw1z Lvyw2 Lxw1w2 .
We will
prove u < (x * y ) * z .
rewrite the current goal using Hue (from left to right).
We will
prove (v * (y * z ) ) + (x * w ) + - (v * w ) < (x * y ) * z .
We will
prove v * (y * z ) + x * w < (x * y ) * z + v * w .
We will
prove (v * (y * z ) + x * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) < ((x * y ) * z + v * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
An exact proof term for the current goal is Lvyzxw .
An exact proof term for the current goal is Lvw1zvyw2xw1w2 .
An exact proof term for the current goal is Lvyzxw1zxyw2vwvw1w2 .
An exact proof term for the current goal is Lxyzvw .
An exact proof term for the current goal is Lvw1zvyw2xw1w2 .
We will
prove (v * (y * z ) + x * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ≤ v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc (v * (y * z ) ) (x * w ) (v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ) Lvyz Lxw Lvw1zvyw2xw1w2 (from right to left).
We will
prove v * (y * z ) + x * w + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ≤ v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
We will
prove x * w + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ≤ x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc (v * (w1 * z ) ) (v * (y * w2 ) ) (x * (w1 * w2 ) ) Lvw1z Lvyw2 Lxw1w2 (from left to right).
rewrite the current goal using
DL v (w1 * z ) (y * w2 ) Hv3 Lw1z Lyw2 (from right to left).
We will
prove x * w + v * (w1 * z + y * w2 ) + x * (w1 * w2 ) ≤ x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
rewrite the current goal using
add_SNo_com_3_0_1 (x * w ) (v * (w1 * z + y * w2 ) ) (x * (w1 * w2 ) ) Lxw Lvw1zyw2 Lxw1w2 (from left to right).
We will
prove v * (w1 * z + y * w2 ) + x * w + x * (w1 * w2 ) ≤ x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
rewrite the current goal using
DL x w (w1 * w2 ) Hx Hw Lw1w2 (from right to left).
We will
prove v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) ≤ x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
rewrite the current goal using
DL v w (w1 * w2 ) Hv3 Hw Lw1w2 (from right to left).
We will
prove v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) ≤ x * (w1 * z ) + x * (y * w2 ) + v * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc (x * (w1 * z ) ) (x * (y * w2 ) ) (v * (w + w1 * w2 ) ) Lxw1z Lxyw2 Lvww1w2 (from left to right).
rewrite the current goal using
DL x (w1 * z ) (y * w2 ) Hx Lw1z Lyw2 (from right to left).
We will
prove v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) ≤ x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) .
An exact proof term for the current goal is H1 .
We will
prove v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) < ((x * y ) * z + v * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
rewrite the current goal using
add_SNo_com ((x * y ) * z ) (v * w ) Lxyz2 Lvw (from left to right).
rewrite the current goal using
add_SNo_assoc (v * w ) ((x * y ) * z ) (v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ) Lvw Lxyz2 Lvw1zvyw2xw1w2 (from right to left).
We will
prove v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) < v * w + (x * y ) * z + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
rewrite the current goal using
add_SNo_rotate_5_2 (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) ) (v * w ) (v * (w1 * w2 ) ) Lvyz Lxw1z Lxyw2 Lvw Lvw1w2 (from left to right).
We will
prove v * w + v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) < v * w + (x * y ) * z + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
We will
prove v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) < (x * y ) * z + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc ((x * y ) * z ) (v * (w1 * z ) ) (v * (y * w2 ) + x * (w1 * w2 ) ) Lxyz2 Lvw1z (SNo_add_SNo (v * (y * w2 ) ) (x * (w1 * w2 ) ) Lvyw2 Lxw1w2 ) (from left to right).
We will
prove v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) < ((x * y ) * z + v * (w1 * z ) ) + v * (y * w2 ) + x * (w1 * w2 ) .
rewrite the current goal using IHxz v Hv w2 Hw2 (from left to right).
rewrite the current goal using IHyz w1 Hw1 w2 Hw2 (from left to right).
We will
prove v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) < ((x * y ) * z + v * (w1 * z ) ) + (v * y ) * w2 + (x * w1 ) * w2 .
rewrite the current goal using
DR (v * y ) (x * w1 ) w2 Lvy Lxw1 Hw23 (from right to left).
We will
prove v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) < ((x * y ) * z + v * (w1 * z ) ) + (v * y + x * w1 ) * w2 .
rewrite the current goal using IHxy v Hv w1 Hw1 (from left to right).
rewrite the current goal using
DR (x * y ) (v * w1 ) z Lxy Lvw1 Hz (from right to left).
We will
prove v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
rewrite the current goal using
add_SNo_rotate_4_1 (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) ) (v * (w1 * w2 ) ) Lvyz Lxw1z Lxyw2 Lvw1w2 (from right to left).
We will
prove v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * (w1 * w2 ) < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
rewrite the current goal using
add_SNo_assoc (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) + v * (w1 * w2 ) ) Lvyz Lxw1z (SNo_add_SNo (x * (y * w2 ) ) (v * (w1 * w2 ) ) Lxyw2 Lvw1w2 ) (from left to right).
We will
prove (v * (y * z ) + x * (w1 * z ) ) + x * (y * w2 ) + v * (w1 * w2 ) < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
rewrite the current goal using IHx v Hv (from left to right).
rewrite the current goal using IHy w1 Hw1 (from left to right).
rewrite the current goal using IHz w2 Hw2 (from left to right).
rewrite the current goal using IHxyz v Hv w1 Hw1 w2 Hw2 (from left to right).
We will
prove ((v * y ) * z + (x * w1 ) * z ) + (x * y ) * w2 + (v * w1 ) * w2 < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
rewrite the current goal using
DR (v * y ) (x * w1 ) z Lvy Lxw1 Hz (from right to left).
rewrite the current goal using
DR (x * y ) (v * w1 ) w2 Lxy Lvw1 Hw23 (from right to left).
We will
prove (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
An exact proof term for the current goal is H2 .
Apply HLE u Hu to the current goal.
Let v be given.
Let w be given.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
We prove the intermediate
claim Lw :
SNo w .
Apply SNoL_E (y * z ) Lyz w Hw to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
Apply IL y z Hy Hz w Hw to the current goal.
Let w1 be given.
Let w2 be given.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We will
prove v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) ≤ x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) .
Apply M_Le x (w1 * z + y * w2 ) v (w + w1 * w2 ) Hx Lw1zyw2 Hv1 Lww1w2 to the current goal.
An exact proof term for the current goal is Hv3 .
We will
prove w + w1 * w2 ≤ w1 * z + y * w2 .
An exact proof term for the current goal is Hwl .
We will
prove (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
Apply M_Lt (x * y + v * w1 ) z (v * y + x * w1 ) w2 (SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) ) Hz (SNo_add_SNo (v * y ) (x * w1 ) (SNo_M v y Hv1 Hy ) (SNo_M x w1 Hx Hw11 ) ) Hw21 to the current goal.
We will
prove v * y + x * w1 < x * y + v * w1 .
Apply M_Lt x y v w1 Hx Hy Hv1 Hw11 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
Let w1 be given.
Let w2 be given.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv1 Hy .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11 .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw11 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We will
prove v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) ≤ x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) .
Apply M_Le x (w1 * z + y * w2 ) v (w + w1 * w2 ) Hx Lw1zyw2 Hv1 Lww1w2 to the current goal.
An exact proof term for the current goal is Hv3 .
We will
prove w + w1 * w2 ≤ w1 * z + y * w2 .
An exact proof term for the current goal is Hwl .
We will
prove (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) (SNo_M v y Hv1 Hy ) (SNo_M x w1 Hx Hw11 ) .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) .
We prove the intermediate
claim Lvyxw1z :
SNo ((v * y + x * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) z Lvyxw1 Hz .
We prove the intermediate
claim Lxyvw1z :
SNo ((x * y + v * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) z Lxyvw1 Hz .
We prove the intermediate
claim Lvyxw1w2 :
SNo ((v * y + x * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) w2 Lvyxw1 Hw21 .
We prove the intermediate
claim Lxyvw1w2 :
SNo ((x * y + v * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) w2 Lxyvw1 Hw21 .
rewrite the current goal using
add_SNo_com ((v * y + x * w1 ) * z ) ((x * y + v * w1 ) * w2 ) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using
add_SNo_com ((x * y + v * w1 ) * z ) ((v * y + x * w1 ) * w2 ) Lxyvw1z Lvyxw1w2 (from left to right).
We will
prove (x * y + v * w1 ) * w2 + (v * y + x * w1 ) * z < (v * y + x * w1 ) * w2 + (x * y + v * w1 ) * z .
Apply M_Lt (v * y + x * w1 ) w2 (x * y + v * w1 ) z Lvyxw1 Hw21 Lxyvw1 Hz to the current goal.
We will
prove x * y + v * w1 < v * y + x * w1 .
rewrite the current goal using
add_SNo_com (x * y ) (v * w1 ) Lxy Lvw1 (from left to right).
rewrite the current goal using
add_SNo_com (v * y ) (x * w1 ) Lvy Lxw1 (from left to right).
Apply M_Lt x w1 v y Hx Hw11 Hv1 Hy to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
Let v be given.
Let w be given.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
We prove the intermediate
claim Lw :
SNo w .
Apply SNoR_E (y * z ) Lyz w Hw to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An exact proof term for the current goal is SNo_M v w Hv1 Lw .
Apply IR y z Hy Hz w Hw to the current goal.
Let w1 be given.
Let w2 be given.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv1 Hy .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11 .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw11 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We prove the intermediate
claim Lxww1w2 :
SNo (x * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w + w1 * w2 ) Hx Lww1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv1 Lww1w2 .
We prove the intermediate
claim Lvw1w2 :
SNo (v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * w2 ) Hv1 Lw1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv1 Lw1zyw2 .
We prove the intermediate
claim Lvwvw1w2 :
SNo (v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (v * (w1 * w2 ) ) Lvw Lvw1w2 .
We prove the intermediate
claim Lxw1zyw2 :
SNo (x * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z + y * w2 ) Hx Lw1zyw2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) Lvy Lxw1 .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy Lvw1 .
We prove the intermediate
claim Lvyxw1z :
SNo ((v * y + x * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) z Lvyxw1 Hz .
We prove the intermediate
claim Lxyvw1z :
SNo ((x * y + v * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) z Lxyvw1 Hz .
We prove the intermediate
claim Lvyxw1w2 :
SNo ((v * y + x * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) w2 Lvyxw1 Hw21 .
We prove the intermediate
claim Lxyvw1w2 :
SNo ((x * y + v * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) w2 Lxyvw1 Hw21 .
We will
prove v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) ≤ x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_com (v * (w1 * z + y * w2 ) ) (x * (w + w1 * w2 ) ) Lvw1zyw2 Lxww1w2 (from left to right).
We will
prove x * (w + w1 * w2 ) + v * (w1 * z + y * w2 ) ≤ x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_com (x * (w1 * z + y * w2 ) ) (v * (w + w1 * w2 ) ) Lxw1zyw2 Lvww1w2 (from left to right).
We will
prove x * (w + w1 * w2 ) + v * (w1 * z + y * w2 ) ≤ v * (w + w1 * w2 ) + x * (w1 * z + y * w2 ) .
Apply M_Le v (w + w1 * w2 ) x (w1 * z + y * w2 ) Hv1 Lww1w2 Hx Lw1zyw2 to the current goal.
An exact proof term for the current goal is Hv3 .
We will
prove w1 * z + y * w2 ≤ w + w1 * w2 .
An exact proof term for the current goal is Hwl .
We will
prove (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
rewrite the current goal using
add_SNo_com ((v * y + x * w1 ) * z ) ((x * y + v * w1 ) * w2 ) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using
add_SNo_com ((x * y + v * w1 ) * z ) ((v * y + x * w1 ) * w2 ) Lxyvw1z Lvyxw1w2 (from left to right).
We will
prove (x * y + v * w1 ) * w2 + (v * y + x * w1 ) * z < (v * y + x * w1 ) * w2 + (x * y + v * w1 ) * z .
Apply M_Lt (v * y + x * w1 ) w2 (x * y + v * w1 ) z Lvyxw1 Hw21 Lxyvw1 Hz to the current goal.
We will
prove x * y + v * w1 < v * y + x * w1 .
Apply M_Lt v y x w1 Hv1 Hy Hx Hw11 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
Let w1 be given.
Let w2 be given.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv1 Hy .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11 .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw11 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We prove the intermediate
claim Lxww1w2 :
SNo (x * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w + w1 * w2 ) Hx Lww1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv1 Lww1w2 .
We prove the intermediate
claim Lvw1w2 :
SNo (v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * w2 ) Hv1 Lw1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv1 Lw1zyw2 .
We prove the intermediate
claim Lvwvw1w2 :
SNo (v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (v * (w1 * w2 ) ) Lvw Lvw1w2 .
We prove the intermediate
claim Lxw1zyw2 :
SNo (x * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z + y * w2 ) Hx Lw1zyw2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) Lvy Lxw1 .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy Lvw1 .
We prove the intermediate
claim Lvyxw1z :
SNo ((v * y + x * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) z Lvyxw1 Hz .
We prove the intermediate
claim Lxyvw1z :
SNo ((x * y + v * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) z Lxyvw1 Hz .
We prove the intermediate
claim Lvyxw1w2 :
SNo ((v * y + x * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) w2 Lvyxw1 Hw21 .
We prove the intermediate
claim Lxyvw1w2 :
SNo ((x * y + v * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) w2 Lxyvw1 Hw21 .
We will
prove v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) ≤ x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_com (v * (w1 * z + y * w2 ) ) (x * (w + w1 * w2 ) ) Lvw1zyw2 Lxww1w2 (from left to right).
We will
prove x * (w + w1 * w2 ) + v * (w1 * z + y * w2 ) ≤ x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_com (x * (w1 * z + y * w2 ) ) (v * (w + w1 * w2 ) ) Lxw1zyw2 Lvww1w2 (from left to right).
We will
prove x * (w + w1 * w2 ) + v * (w1 * z + y * w2 ) ≤ v * (w + w1 * w2 ) + x * (w1 * z + y * w2 ) .
Apply M_Le v (w + w1 * w2 ) x (w1 * z + y * w2 ) Hv1 Lww1w2 Hx Lw1zyw2 to the current goal.
An exact proof term for the current goal is Hv3 .
We will
prove w1 * z + y * w2 ≤ w + w1 * w2 .
An exact proof term for the current goal is Hwl .
We will
prove (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 < (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) (SNo_M v y Hv1 Hy ) (SNo_M x w1 Hx Hw11 ) .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) .
Apply M_Lt (x * y + v * w1 ) z (v * y + x * w1 ) w2 Lxyvw1 Hz Lvyxw1 Hw21 to the current goal.
We will
prove v * y + x * w1 < x * y + v * w1 .
rewrite the current goal using
add_SNo_com (x * y ) (v * w1 ) Lxy Lvw1 (from left to right).
rewrite the current goal using
add_SNo_com (v * y ) (x * w1 ) Lvy Lxw1 (from left to right).
We will
prove x * w1 + v * y < v * w1 + x * y .
Apply M_Lt v w1 x y Hv1 Hw11 Hx Hy to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
∎
Theorem. (
mul_SNo_assoc_lem2 ) The following is provable:
∀x y z, SNo x → SNo y → SNo z → (∀ u ∈ SNoS_ (SNoLev x ) , u * (y * z ) = (u * y ) * z ) → (∀ v ∈ SNoS_ (SNoLev y ) , x * (v * z ) = (x * v ) * z ) → (∀ w ∈ SNoS_ (SNoLev z ) , x * (y * w ) = (x * y ) * w ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ v ∈ SNoS_ (SNoLev y ) , u * (v * z ) = (u * v ) * z ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ w ∈ SNoS_ (SNoLev z ) , u * (y * w ) = (u * y ) * w ) → (∀ v ∈ SNoS_ (SNoLev y ) , ∀ w ∈ SNoS_ (SNoLev z ) , x * (v * w ) = (x * v ) * w ) → (∀ u ∈ SNoS_ (SNoLev x ) , ∀ v ∈ SNoS_ (SNoLev y ) , ∀ w ∈ SNoS_ (SNoLev z ) , u * (v * w ) = (u * v ) * w ) → ∀R, (∀ u ∈ R , ∀q : prop , (∀ v ∈ SNoL x , ∀ w ∈ SNoR (y * z ) , u = v * (y * z ) + x * w + - v * w → q ) → (∀ v ∈ SNoR x , ∀ w ∈ SNoL (y * z ) , u = v * (y * z ) + x * w + - v * w → q ) → q ) → ∀ u ∈ R , (x * y ) * z < u
Proof: Let x, y and z be given.
Assume Hx Hy Hz .
Assume IHx IHy IHz IHxy IHxz IHyz IHxyz .
Let R be given.
Assume HRE .
Let u be given.
We will
prove (x * y ) * z < u .
We prove the intermediate
claim Lxy :
SNo (x * y ) .
An exact proof term for the current goal is SNo_M x y Hx Hy .
We prove the intermediate
claim Lyz :
SNo (y * z ) .
An exact proof term for the current goal is SNo_M y z Hy Hz .
We prove the intermediate
claim Lxyz2 :
SNo ((x * y ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y ) z Lxy Hz .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
Let w1 be given.
Assume Hw1 .
Let w2 be given.
Assume Hw2 .
Assume Hue H1 H2 .
Assume Hv1 Hv2 Hv3 Hv4 .
Assume Hw11 Hw12 Hw13 Hw14 .
Assume Hw21 Hw22 Hw23 Hw24 .
We prove the intermediate
claim Lvyz :
SNo (v * (y * z ) ) .
An
exact proof term for the current goal is
SNo_M v (y * z ) Hv3 Lyz .
We prove the intermediate
claim Lxw :
SNo (x * w ) .
An exact proof term for the current goal is SNo_M x w Hx Hw .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw13 .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An exact proof term for the current goal is SNo_M v w Hv3 Hw .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv3 Hw13 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv3 Hy .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw13 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw23 .
We prove the intermediate
claim Lvw1z :
SNo (v * (w1 * z ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z ) Hv3 Lw1z .
We prove the intermediate
claim Lvyw2 :
SNo (v * (y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (y * w2 ) Hv3 Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw13 Hw23 .
We prove the intermediate
claim Lxw1w2 :
SNo (x * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * w2 ) Hx Lw1w2 .
We prove the intermediate
claim Lxw1z :
SNo (x * (w1 * z ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z ) Hx Lw1z .
We prove the intermediate
claim Lxyw2 :
SNo (x * (y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (y * w2 ) Hx Lyw2 .
We prove the intermediate
claim Lvyzxw :
SNo (v * (y * z ) + x * w ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * (y * z ) ) (x * w ) Lvyz Lxw .
We prove the intermediate
claim Lxyzvw :
SNo ((x * y ) * z + v * w ) .
An
exact proof term for the current goal is
SNo_add_SNo ((x * y ) * z ) (v * w ) Lxyz2 Lvw .
We prove the intermediate
claim Lvw1w2 :
SNo (v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * w2 ) Hv3 Lw1w2 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Hw Lw1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv3 Lww1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv3 (SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 ) .
We prove the intermediate
claim Lvwvw1w2 :
SNo (v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (v * (w1 * w2 ) ) Lvw Lvw1w2 .
We prove the intermediate
claim Lvyzxw1zxyw2vwvw1w2 :
SNo (v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo_4 (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) ) (v * w + v * (w1 * w2 ) ) Lvyz Lxw1z Lxyw2 Lvwvw1w2 .
We prove the intermediate
claim Lvw1zvyw2xw1w2 :
SNo (v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo_3 (v * (w1 * z ) ) (v * (y * w2 ) ) (x * (w1 * w2 ) ) Lvw1z Lvyw2 Lxw1w2 .
rewrite the current goal using Hue (from left to right).
We will
prove (x * y ) * z + v * w < v * (y * z ) + x * w .
We will
prove ((x * y ) * z + v * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) < (v * (y * z ) + x * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
An exact proof term for the current goal is Lxyzvw .
An exact proof term for the current goal is Lvw1zvyw2xw1w2 .
An exact proof term for the current goal is Lvyzxw1zxyw2vwvw1w2 .
An exact proof term for the current goal is Lvyzxw .
An exact proof term for the current goal is Lvw1zvyw2xw1w2 .
We will
prove ((x * y ) * z + v * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) < v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
rewrite the current goal using
add_SNo_com ((x * y ) * z ) (v * w ) Lxyz2 Lvw (from left to right).
rewrite the current goal using
add_SNo_assoc (v * w ) ((x * y ) * z ) (v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ) Lvw Lxyz2 Lvw1zvyw2xw1w2 (from right to left).
We will
prove v * w + (x * y ) * z + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) < v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) .
rewrite the current goal using
add_SNo_rotate_5_2 (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) ) (v * w ) (v * (w1 * w2 ) ) Lvyz Lxw1z Lxyw2 Lvw Lvw1w2 (from left to right).
We will
prove v * w + (x * y ) * z + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) < v * w + v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) .
We will
prove (x * y ) * z + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) < v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) .
rewrite the current goal using
add_SNo_assoc ((x * y ) * z ) (v * (w1 * z ) ) (v * (y * w2 ) + x * (w1 * w2 ) ) Lxyz2 Lvw1z (SNo_add_SNo (v * (y * w2 ) ) (x * (w1 * w2 ) ) Lvyw2 Lxw1w2 ) (from left to right).
We will
prove ((x * y ) * z + v * (w1 * z ) ) + v * (y * w2 ) + x * (w1 * w2 ) < v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) .
rewrite the current goal using IHxz v Hv w2 Hw2 (from left to right).
rewrite the current goal using IHyz w1 Hw1 w2 Hw2 (from left to right).
We will
prove ((x * y ) * z + v * (w1 * z ) ) + (v * y ) * w2 + (x * w1 ) * w2 < v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) .
rewrite the current goal using
DR (v * y ) (x * w1 ) w2 Lvy Lxw1 Hw23 (from right to left).
We will
prove ((x * y ) * z + v * (w1 * z ) ) + (v * y + x * w1 ) * w2 < v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) .
rewrite the current goal using IHxy v Hv w1 Hw1 (from left to right).
rewrite the current goal using
DR (x * y ) (v * w1 ) z Lxy Lvw1 Hz (from right to left).
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < v * (w1 * w2 ) + v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) .
rewrite the current goal using
add_SNo_rotate_4_1 (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) ) (v * (w1 * w2 ) ) Lvyz Lxw1z Lxyw2 Lvw1w2 (from right to left).
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * (w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc (v * (y * z ) ) (x * (w1 * z ) ) (x * (y * w2 ) + v * (w1 * w2 ) ) Lvyz Lxw1z (SNo_add_SNo (x * (y * w2 ) ) (v * (w1 * w2 ) ) Lxyw2 Lvw1w2 ) (from left to right).
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < (v * (y * z ) + x * (w1 * z ) ) + x * (y * w2 ) + v * (w1 * w2 ) .
rewrite the current goal using IHx v Hv (from left to right).
rewrite the current goal using IHy w1 Hw1 (from left to right).
rewrite the current goal using IHz w2 Hw2 (from left to right).
rewrite the current goal using IHxyz v Hv w1 Hw1 w2 Hw2 (from left to right).
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < ((v * y ) * z + (x * w1 ) * z ) + (x * y ) * w2 + (v * w1 ) * w2 .
rewrite the current goal using
DR (v * y ) (x * w1 ) z Lvy Lxw1 Hz (from right to left).
rewrite the current goal using
DR (x * y ) (v * w1 ) w2 Lxy Lvw1 Hw23 (from right to left).
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 .
An exact proof term for the current goal is H2 .
We will
prove v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ≤ (v * (y * z ) + x * w ) + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc (v * (y * z ) ) (x * w ) (v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) ) Lvyz Lxw Lvw1zvyw2xw1w2 (from right to left).
We will
prove v * (y * z ) + x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ≤ v * (y * z ) + x * w + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
We will
prove x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ≤ x * w + v * (w1 * z ) + v * (y * w2 ) + x * (w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc (v * (w1 * z ) ) (v * (y * w2 ) ) (x * (w1 * w2 ) ) Lvw1z Lvyw2 Lxw1w2 (from left to right).
rewrite the current goal using
DL v (w1 * z ) (y * w2 ) Hv3 Lw1z Lyw2 (from right to left).
We will
prove x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ≤ x * w + v * (w1 * z + y * w2 ) + x * (w1 * w2 ) .
rewrite the current goal using
add_SNo_com_3_0_1 (x * w ) (v * (w1 * z + y * w2 ) ) (x * (w1 * w2 ) ) Lxw Lvw1zyw2 Lxw1w2 (from left to right).
We will
prove x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * w + x * (w1 * w2 ) .
rewrite the current goal using
DL x w (w1 * w2 ) Hx Hw Lw1w2 (from right to left).
We will
prove x * (w1 * z ) + x * (y * w2 ) + v * w + v * (w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) .
rewrite the current goal using
DL v w (w1 * w2 ) Hv3 Hw Lw1w2 (from right to left).
We will
prove x * (w1 * z ) + x * (y * w2 ) + v * (w + w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_assoc (x * (w1 * z ) ) (x * (y * w2 ) ) (v * (w + w1 * w2 ) ) Lxw1z Lxyw2 Lvww1w2 (from left to right).
rewrite the current goal using
DL x (w1 * z ) (y * w2 ) Hx Lw1z Lyw2 (from right to left).
We will
prove x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) .
An exact proof term for the current goal is H1 .
Apply HRE u Hu to the current goal.
Let v be given.
Let w be given.
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
We prove the intermediate
claim Lw :
SNo w .
Apply SNoR_E (y * z ) Lyz w Hw to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An exact proof term for the current goal is SNo_M v w Hv1 Lw .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv1 Hy .
Apply IR y z Hy Hz w Hw to the current goal.
Let w1 be given.
Let w2 be given.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw11 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We prove the intermediate
claim Lxww1w2 :
SNo (x * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w + w1 * w2 ) Hx Lww1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv1 Lww1w2 .
We prove the intermediate
claim Lvw1w2 :
SNo (v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * w2 ) Hv1 Lw1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv1 Lw1zyw2 .
We prove the intermediate
claim Lvwvw1w2 :
SNo (v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (v * (w1 * w2 ) ) Lvw Lvw1w2 .
We prove the intermediate
claim Lxw1zyw2 :
SNo (x * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z + y * w2 ) Hx Lw1zyw2 .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) Lvy Lxw1 .
We prove the intermediate
claim Lxyvw1z :
SNo ((x * y + v * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) z Lxyvw1 Hz .
We prove the intermediate
claim Lvyxw1z :
SNo ((v * y + x * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) z Lvyxw1 Hz .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) .
We prove the intermediate
claim Lxyvw1w2 :
SNo ((x * y + v * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) w2 Lxyvw1 Hw21 .
We prove the intermediate
claim Lvyxw1w2 :
SNo ((v * y + x * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) w2 Lvyxw1 Hw21 .
We will
prove x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_com (v * (w1 * z + y * w2 ) ) (x * (w + w1 * w2 ) ) Lvw1zyw2 Lxww1w2 (from left to right).
rewrite the current goal using
add_SNo_com (x * (w1 * z + y * w2 ) ) (v * (w + w1 * w2 ) ) Lxw1zyw2 Lvww1w2 (from left to right).
We will
prove v * (w + w1 * w2 ) + x * (w1 * z + y * w2 ) ≤ x * (w + w1 * w2 ) + v * (w1 * z + y * w2 ) .
Apply M_Le x (w + w1 * w2 ) v (w1 * z + y * w2 ) Hx Lww1w2 Hv1 Lw1zyw2 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hwl .
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 .
rewrite the current goal using
add_SNo_com ((v * y + x * w1 ) * z ) ((x * y + v * w1 ) * w2 ) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using
add_SNo_com ((x * y + v * w1 ) * z ) ((v * y + x * w1 ) * w2 ) Lxyvw1z Lvyxw1w2 (from left to right).
We will
prove (v * y + x * w1 ) * w2 + (x * y + v * w1 ) * z < (x * y + v * w1 ) * w2 + (v * y + x * w1 ) * z .
Apply M_Lt (x * y + v * w1 ) w2 (v * y + x * w1 ) z (SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) ) Hw21 (SNo_add_SNo (v * y ) (x * w1 ) (SNo_M v y Hv1 Hy ) (SNo_M x w1 Hx Hw11 ) ) Hz to the current goal.
We will
prove v * y + x * w1 < x * y + v * w1 .
Apply M_Lt x y v w1 Hx Hy Hv1 Hw11 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
Let w1 be given.
Let w2 be given.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv1 Hy .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11 .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw11 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv1 Lww1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv1 Lw1zyw2 .
We prove the intermediate
claim Lxww1w2 :
SNo (x * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w + w1 * w2 ) Hx Lww1w2 .
We prove the intermediate
claim Lxw1zyw2 :
SNo (x * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z + y * w2 ) Hx Lw1zyw2 .
We will
prove x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) .
rewrite the current goal using
add_SNo_com (v * (w1 * z + y * w2 ) ) (x * (w + w1 * w2 ) ) Lvw1zyw2 Lxww1w2 (from left to right).
rewrite the current goal using
add_SNo_com (x * (w1 * z + y * w2 ) ) (v * (w + w1 * w2 ) ) Lxw1zyw2 Lvww1w2 (from left to right).
Apply M_Le x (w + w1 * w2 ) v (w1 * z + y * w2 ) Hx Lww1w2 Hv1 Lw1zyw2 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hwl .
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) (SNo_M v y Hv1 Hy ) (SNo_M x w1 Hx Hw11 ) .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) .
Apply M_Lt (v * y + x * w1 ) z (x * y + v * w1 ) w2 Lvyxw1 Hz Lxyvw1 Hw21 to the current goal.
We will
prove x * y + v * w1 < v * y + x * w1 .
rewrite the current goal using
add_SNo_com (x * y ) (v * w1 ) Lxy Lvw1 (from left to right).
rewrite the current goal using
add_SNo_com (v * y ) (x * w1 ) Lvy Lxw1 (from left to right).
Apply M_Lt x w1 v y Hx Hw11 Hv1 Hy to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
Let v be given.
Let w be given.
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 Hv2 Hv3 .
We prove the intermediate
claim Lw :
SNo w .
Apply SNoL_E (y * z ) Lyz w Hw to the current goal.
Assume H _ _ .
An exact proof term for the current goal is H .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An exact proof term for the current goal is SNo_M v w Hv1 Lw .
Apply IL y z Hy Hz w Hw to the current goal.
Let w1 be given.
Let w2 be given.
Apply SNoL_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoL_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv1 Hy .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11 .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw11 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We prove the intermediate
claim Lxww1w2 :
SNo (x * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w + w1 * w2 ) Hx Lww1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv1 Lww1w2 .
We prove the intermediate
claim Lvw1w2 :
SNo (v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * w2 ) Hv1 Lw1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv1 Lw1zyw2 .
We prove the intermediate
claim Lvwvw1w2 :
SNo (v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (v * (w1 * w2 ) ) Lvw Lvw1w2 .
We prove the intermediate
claim Lxw1zyw2 :
SNo (x * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z + y * w2 ) Hx Lw1zyw2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) Lvy Lxw1 .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy Lvw1 .
We prove the intermediate
claim Lvyxw1z :
SNo ((v * y + x * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) z Lvyxw1 Hz .
We prove the intermediate
claim Lxyvw1z :
SNo ((x * y + v * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) z Lxyvw1 Hz .
We prove the intermediate
claim Lvyxw1w2 :
SNo ((v * y + x * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) w2 Lvyxw1 Hw21 .
We prove the intermediate
claim Lxyvw1w2 :
SNo ((x * y + v * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) w2 Lxyvw1 Hw21 .
We will
prove x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) .
Apply M_Le v (w1 * z + y * w2 ) x (w + w1 * w2 ) Hv1 Lw1zyw2 Hx Lww1w2 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hwl .
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 .
Apply M_Lt (v * y + x * w1 ) z (x * y + v * w1 ) w2 Lvyxw1 Hz Lxyvw1 Hw21 to the current goal.
We will
prove x * y + v * w1 < v * y + x * w1 .
Apply M_Lt v y x w1 Hv1 Hy Hx Hw11 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
Let w1 be given.
Let w2 be given.
Apply SNoR_E y Hy w1 Hw1 to the current goal.
Assume Hw11 Hw12 Hw13 .
Apply SNoR_E z Hz w2 Hw2 to the current goal.
Assume Hw21 Hw22 Hw23 .
We prove the intermediate
claim Lvy :
SNo (v * y ) .
An exact proof term for the current goal is SNo_M v y Hv1 Hy .
We prove the intermediate
claim Lvw1 :
SNo (v * w1 ) .
An exact proof term for the current goal is SNo_M v w1 Hv1 Hw11 .
We prove the intermediate
claim Lxw1 :
SNo (x * w1 ) .
An exact proof term for the current goal is SNo_M x w1 Hx Hw11 .
We prove the intermediate
claim Lw1z :
SNo (w1 * z ) .
An exact proof term for the current goal is SNo_M w1 z Hw11 Hz .
We prove the intermediate
claim Lyw2 :
SNo (y * w2 ) .
An exact proof term for the current goal is SNo_M y w2 Hy Hw21 .
We prove the intermediate
claim Lw1zyw2 :
SNo (w1 * z + y * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo (w1 * z ) (y * w2 ) Lw1z Lyw2 .
We prove the intermediate
claim Lw1w2 :
SNo (w1 * w2 ) .
An exact proof term for the current goal is SNo_M w1 w2 Hw11 Hw21 .
We prove the intermediate
claim Lww1w2 :
SNo (w + w1 * w2 ) .
An
exact proof term for the current goal is
SNo_add_SNo w (w1 * w2 ) Lw Lw1w2 .
We prove the intermediate
claim Lxww1w2 :
SNo (x * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w + w1 * w2 ) Hx Lww1w2 .
We prove the intermediate
claim Lvww1w2 :
SNo (v * (w + w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w + w1 * w2 ) Hv1 Lww1w2 .
We prove the intermediate
claim Lvw1w2 :
SNo (v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * w2 ) Hv1 Lw1w2 .
We prove the intermediate
claim Lvw1zyw2 :
SNo (v * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M v (w1 * z + y * w2 ) Hv1 Lw1zyw2 .
We prove the intermediate
claim Lvwvw1w2 :
SNo (v * w + v * (w1 * w2 ) ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * w ) (v * (w1 * w2 ) ) Lvw Lvw1w2 .
We prove the intermediate
claim Lxw1zyw2 :
SNo (x * (w1 * z + y * w2 ) ) .
An
exact proof term for the current goal is
SNo_M x (w1 * z + y * w2 ) Hx Lw1zyw2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) Lvy Lxw1 .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy Lvw1 .
We prove the intermediate
claim Lvyxw1z :
SNo ((v * y + x * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) z Lvyxw1 Hz .
We prove the intermediate
claim Lxyvw1z :
SNo ((x * y + v * w1 ) * z ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) z Lxyvw1 Hz .
We prove the intermediate
claim Lvyxw1w2 :
SNo ((v * y + x * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (v * y + x * w1 ) w2 Lvyxw1 Hw21 .
We prove the intermediate
claim Lxyvw1w2 :
SNo ((x * y + v * w1 ) * w2 ) .
An
exact proof term for the current goal is
SNo_M (x * y + v * w1 ) w2 Lxyvw1 Hw21 .
We will
prove x * (w1 * z + y * w2 ) + v * (w + w1 * w2 ) ≤ v * (w1 * z + y * w2 ) + x * (w + w1 * w2 ) .
Apply M_Le v (w1 * z + y * w2 ) x (w + w1 * w2 ) Hv1 Lw1zyw2 Hx Lww1w2 to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hwl .
We will
prove (x * y + v * w1 ) * z + (v * y + x * w1 ) * w2 < (v * y + x * w1 ) * z + (x * y + v * w1 ) * w2 .
We prove the intermediate
claim Lvyxw1 :
SNo (v * y + x * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (v * y ) (x * w1 ) (SNo_M v y Hv1 Hy ) (SNo_M x w1 Hx Hw11 ) .
We prove the intermediate
claim Lxyvw1 :
SNo (x * y + v * w1 ) .
An
exact proof term for the current goal is
SNo_add_SNo (x * y ) (v * w1 ) Lxy (SNo_M v w1 Hv1 Hw11 ) .
rewrite the current goal using
add_SNo_com ((v * y + x * w1 ) * z ) ((x * y + v * w1 ) * w2 ) Lvyxw1z Lxyvw1w2 (from left to right).
rewrite the current goal using
add_SNo_com ((x * y + v * w1 ) * z ) ((v * y + x * w1 ) * w2 ) Lxyvw1z Lvyxw1w2 (from left to right).
We will
prove (v * y + x * w1 ) * w2 + (x * y + v * w1 ) * z < (x * y + v * w1 ) * w2 + (v * y + x * w1 ) * z .
Apply M_Lt (x * y + v * w1 ) w2 (v * y + x * w1 ) z Lxyvw1 Hw21 Lvyxw1 Hz to the current goal.
We will
prove v * y + x * w1 < x * y + v * w1 .
rewrite the current goal using
add_SNo_com (x * y ) (v * w1 ) Lxy Lvw1 (from left to right).
rewrite the current goal using
add_SNo_com (v * y ) (x * w1 ) Lvy Lxw1 (from left to right).
We will
prove x * w1 + v * y < v * w1 + x * y .
Apply M_Lt v w1 x y Hv1 Hw11 Hx Hy to the current goal.
An exact proof term for the current goal is Hv3 .
An exact proof term for the current goal is Hw13 .
An exact proof term for the current goal is Hw23 .
∎
End of Section mul_SNo_assoc_lems
Proof: Set P to be the term
λx y z ⇒ x * (y * z ) = (x * y ) * z of type
set → set → set → prop .
We will
prove ∀x y z, SNo x → SNo y → SNo z → P x y z .
Let x, y and z be given.
Assume Hx Hy Hz .
We will
prove x * (y * z ) = (x * y ) * z .
We prove the intermediate
claim Lxy :
SNo (x * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo x y Hx Hy .
We prove the intermediate
claim Lyz :
SNo (y * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo y z Hy Hz .
We prove the intermediate
claim Lxyz1 :
SNo (x * (y * z ) ) .
An
exact proof term for the current goal is
SNo_mul_SNo x (y * z ) Hx Lyz .
We prove the intermediate
claim Lxyz2 :
SNo ((x * y ) * z ) .
An
exact proof term for the current goal is
SNo_mul_SNo (x * y ) z Lxy Hz .
Let L and R be given.
Assume HLR HLE HLI1 HLI2 HRE HRI1 HRI2 .
Let L' and R' be given.
Assume HLR' HLE' HLI1' HLI2' HRE' HRI1' HRI2' .
rewrite the current goal using HE (from left to right).
rewrite the current goal using HE' (from left to right).
Let x and y be given.
Assume Hx Hy .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp1 Hp2 .
Let w be given.
Let v be given.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 _ _ .
We will
prove u + w * v ≤ w * x + y * v → p .
An exact proof term for the current goal is Hp1 v Hv w Hw .
Let w be given.
Let v be given.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 _ _ .
We will
prove u + w * v ≤ w * x + y * v → p .
An exact proof term for the current goal is Hp2 v Hv w Hw .
Let x and y be given.
Assume Hx Hy .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp1 Hp2 .
Let w be given.
Let v be given.
Apply SNoL_E y Hy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoR_E x Hx v Hv to the current goal.
Assume Hv1 _ _ .
An exact proof term for the current goal is Hp2 v Hv w Hw .
Let w be given.
Let v be given.
Apply SNoR_E y Hy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoL_E x Hx v Hv to the current goal.
Assume Hv1 _ _ .
An exact proof term for the current goal is Hp1 v Hv w Hw .
We prove the intermediate
claim LMLt' :
∀x y u v, SNo x → SNo y → SNo u → SNo v → u < x → v < y → y * u + v * x < y * x + v * u .
Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy .
We will
prove v * x + y * u < y * x + v * u .
An
exact proof term for the current goal is
mul_SNo_Lt y x v u Hy Hx Hv Hu Hvy Hux .
We prove the intermediate
claim LMLe' :
∀x y u v, SNo x → SNo y → SNo u → SNo v → u ≤ x → v ≤ y → y * u + v * x ≤ y * x + v * u .
Let x, y, u and v be given.
Assume Hx Hy Hu Hv Hux Hvy .
We will
prove v * x + y * u ≤ y * x + v * u .
An
exact proof term for the current goal is
mul_SNo_Le y x v u Hy Hx Hv Hu Hvy Hux .
Let u be given.
Assume Hu .
Use symmetry.
An exact proof term for the current goal is IHx u Hu .
Let v be given.
Assume Hv .
Use symmetry.
An exact proof term for the current goal is IHy v Hv .
Let w be given.
Assume Hw .
Use symmetry.
An exact proof term for the current goal is IHz w Hw .
Let v be given.
Assume Hv .
Let u be given.
Assume Hu .
Use symmetry.
An exact proof term for the current goal is IHxy u Hu v Hv .
Let w be given.
Assume Hw .
Let u be given.
Assume Hu .
Use symmetry.
An exact proof term for the current goal is IHxz u Hu w Hw .
Let w be given.
Assume Hw .
Let v be given.
Assume Hv .
Use symmetry.
An exact proof term for the current goal is IHyz v Hv w Hw .
Let w be given.
Assume Hw .
Let v be given.
Assume Hv .
Let u be given.
Assume Hu .
Use symmetry.
An exact proof term for the current goal is IHxyz u Hu v Hv w Hw .
An exact proof term for the current goal is HLR .
An exact proof term for the current goal is HLR' .
rewrite the current goal using HE' (from right to left).
We will
prove ∀ u ∈ L , u < (x * y ) * z .
rewrite the current goal using HE' (from right to left).
We will
prove ∀ u ∈ R , (x * y ) * z < u .
rewrite the current goal using HE (from right to left).
We will
prove ∀ u ∈ L' , u < x * (y * z ) .
Let u be given.
Assume Hu .
Let q be given.
Assume Hq1 Hq2 .
Apply HLE' u Hu to the current goal.
Let w be given.
Assume Hw .
Let v be given.
Assume Hv .
Apply SNoL_E (x * y ) Lxy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoL_E z Hz v Hv to the current goal.
Assume Hv1 _ _ .
We will
prove u = w * z + (x * y ) * v + - w * v → q .
An exact proof term for the current goal is Hq1 v Hv w Hw .
Let w be given.
Assume Hw .
Let v be given.
Assume Hv .
Apply SNoR_E (x * y ) Lxy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoR_E z Hz v Hv to the current goal.
Assume Hv1 _ _ .
We will
prove u = w * z + (x * y ) * v + - w * v → q .
An exact proof term for the current goal is Hq2 v Hv w Hw .
rewrite the current goal using HE (from right to left).
We will
prove ∀ u ∈ R' , x * (y * z ) < u .
Let u be given.
Assume Hu .
Let q be given.
Assume Hq1 Hq2 .
Apply HRE' u Hu to the current goal.
Let w be given.
Assume Hw .
Let v be given.
Assume Hv .
Apply SNoL_E (x * y ) Lxy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoR_E z Hz v Hv to the current goal.
Assume Hv1 _ _ .
We will
prove u = w * z + (x * y ) * v + - w * v → q .
An exact proof term for the current goal is Hq2 v Hv w Hw .
Let w be given.
Assume Hw .
Let v be given.
Assume Hv .
Apply SNoR_E (x * y ) Lxy w Hw to the current goal.
Assume Hw1 _ _ .
Apply SNoL_E z Hz v Hv to the current goal.
Assume Hv1 _ _ .
We will
prove u = w * z + (x * y ) * v + - w * v → q .
An exact proof term for the current goal is Hq1 v Hv w Hw .
∎
Proof: Let n be given.
We prove the intermediate
claim Ln1 :
nat_p n .
An
exact proof term for the current goal is
omega_nat_p n Hn .
We prove the intermediate
claim Ln2 :
ordinal n .
We prove the intermediate
claim Ln3 :
SNo n .
An
exact proof term for the current goal is
ordinal_SNo n Ln2 .
We prove the intermediate
claim L1 :
∀m, nat_p m → mul_nat n m = n * m .
rewrite the current goal using
mul_SNo_zeroR n Ln3 (from left to right).
An
exact proof term for the current goal is
mul_nat_0R n .
Let m be given.
An
exact proof term for the current goal is
mul_nat_SR n m Hm .
Use f_equal.
An exact proof term for the current goal is IH .
Use symmetry.
We will
prove n * (1 + m ) = n + n * m .
We will
prove n * 1 + n * m = n + n * m .
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_oneR n Ln3 .
Let m be given.
An
exact proof term for the current goal is
L1 m (omega_nat_p m Hm ) .
∎
Proof: Let n be given.
Assume Hn .
Let m be given.
Assume Hm .
rewrite the current goal using
mul_nat_mul_SNo n Hn m Hm (from right to left).
∎
Proof: Let x be given.
Assume Hx .
∎
Proof: Let x be given.
Assume Hx .
An
exact proof term for the current goal is
mul_SNo_oneR x Hx .
∎
Proof: Let x be given.
Assume Hx Hx1 .
rewrite the current goal using
add_SNo_1_1_2 (from right to left).
rewrite the current goal using
mul_SNo_oneL x Hx (from left to right).
We will
prove x + 1 < x + x .
An exact proof term for the current goal is Hx1 .
∎
Proof: Let x, y and z be given.
Assume Hx1 Hx2 Hy Hz Hyz .
We will
prove x * y < x * z .
We prove the intermediate
claim L1 :
0 * z + x * y = x * y .
Use transitivity with and
0 + x * y .
We prove the intermediate
claim L2 :
x * z + 0 * y = x * z .
Use transitivity with and
x * z + 0 .
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
An
exact proof term for the current goal is
mul_SNo_Lt x z 0 y Hx1 Hz SNo_0 Hy Hx2 Hyz .
∎
Proof: Let x, y and z be given.
Assume Hx1 Hx2 Hy Hz Hyz .
We will
prove x * y ≤ x * z .
We prove the intermediate
claim L1 :
0 * z + x * y = x * y .
Use transitivity with and
0 + x * y .
We prove the intermediate
claim L2 :
x * z + 0 * y = x * z .
Use transitivity with and
x * z + 0 .
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
An
exact proof term for the current goal is
mul_SNo_Le x z 0 y Hx1 Hz SNo_0 Hy Hx2 Hyz .
∎
Proof: Let x, y and z be given.
Assume Hx1 Hx2 Hy Hz Hzy .
We will
prove x * y < x * z .
We prove the intermediate
claim L1 :
x * y + 0 * z = x * y .
Use transitivity with and
x * y + 0 .
We prove the intermediate
claim L2 :
0 * y + x * z = x * z .
Use transitivity with and
0 + x * z .
rewrite the current goal using L1 (from right to left).
rewrite the current goal using L2 (from right to left).
An
exact proof term for the current goal is
mul_SNo_Lt 0 y x z SNo_0 Hy Hx1 Hz Hx2 Hzy .
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz Hzpos Hxy .
rewrite the current goal using
mul_SNo_com x z Hx Hz (from left to right).
rewrite the current goal using
mul_SNo_com y z Hy Hz (from left to right).
An
exact proof term for the current goal is
pos_mul_SNo_Lt z x y Hz Hzpos Hx Hy Hxy .
∎
Proof: Let x and y be given.
Assume Hx Hy Hx1 Hy0 .
rewrite the current goal using
mul_SNo_oneL y Hy (from right to left) at position 2.
We will
prove x * y < 1 * y .
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz Hznn Hxy .
rewrite the current goal using
mul_SNo_com x z Hx Hz (from left to right).
rewrite the current goal using
mul_SNo_com y z Hy Hz (from left to right).
An
exact proof term for the current goal is
nonneg_mul_SNo_Le z x y Hz Hznn Hx Hy Hxy .
∎
Proof: Let x and y be given.
Assume Hx Hy Hx1 Hy0 .
rewrite the current goal using
mul_SNo_oneL y Hy (from right to left) at position 2.
We will
prove x * y ≤ 1 * y .
∎
Proof: Let x, y, z and w be given.
Assume Hx Hy Hz Hw Hxpos Hypos Hxz Hyw .
We will
prove x * y < x * w .
An
exact proof term for the current goal is
pos_mul_SNo_Lt x y w Hx Hxpos Hy Hw Hyw .
We will
prove x * w < z * w .
An
exact proof term for the current goal is
SNoLt_tra 0 y w SNo_0 Hy Hw Hypos Hyw .
An exact proof term for the current goal is Hxz .
∎
Proof: Let x, y, z and w be given.
Assume Hx Hy Hz Hw Hxnn Hynn Hxz Hyw .
We will
prove x * y ≤ x * w .
An
exact proof term for the current goal is
nonneg_mul_SNo_Le x y w Hx Hxnn Hy Hw Hyw .
We will
prove x * w ≤ z * w .
An
exact proof term for the current goal is
SNoLe_tra 0 y w SNo_0 Hy Hw Hynn Hyw .
An exact proof term for the current goal is Hxz .
∎
Proof: Let x and y be given.
Assume Hx Hy Hx0 Hy0 .
We will
prove 0 + 0 < x * y .
rewrite the current goal using
mul_SNo_zeroR x Hx (from right to left) at position 2.
rewrite the current goal using
mul_SNo_zeroR y Hy (from right to left) at position 1.
∎
Proof: Let x and y be given.
Assume Hx Hy Hx0 Hy0 .
rewrite the current goal using
mul_SNo_zeroR y Hy (from right to left) at position 3.
rewrite the current goal using
mul_SNo_zeroR x Hx (from right to left) at position 2.
∎
Proof: Let x and y be given.
Assume Hx Hy Hx0 Hy0 .
rewrite the current goal using
mul_SNo_zeroR x Hx (from right to left) at position 3.
rewrite the current goal using
mul_SNo_zeroR y Hy (from right to left) at position 2.
∎
Proof: Let x and y be given.
Assume Hx Hy Hx0 Hy0 .
We will
prove 0 + 0 < x * y .
rewrite the current goal using
mul_SNo_zeroR y Hy (from right to left) at position 2.
rewrite the current goal using
mul_SNo_zeroR x Hx (from right to left) at position 1.
∎
Proof: Let x and y be given.
Assume Hx Hy Hxnn Hynn .
An
exact proof term for the current goal is
mul_SNo_pos_pos x y Hx Hy H1 H2 .
rewrite the current goal using H2 (from right to left).
rewrite the current goal using
mul_SNo_zeroR x Hx (from left to right).
rewrite the current goal using H1 (from right to left).
rewrite the current goal using
mul_SNo_zeroL y Hy (from left to right).
∎
Proof: Let x and y be given.
Assume Hx Hy Hxnp Hypos .
An
exact proof term for the current goal is
mul_SNo_neg_pos x y Hx Hy H1 Hypos .
We prove the intermediate
claim L1 :
x = 0 .
An exact proof term for the current goal is Hxnp .
An exact proof term for the current goal is H1 .
rewrite the current goal using L1 (from left to right).
rewrite the current goal using
mul_SNo_zeroL y Hy (from left to right).
∎
Proof: Let x and y be given.
Assume Hx Hy Hxnp Hyneg .
An
exact proof term for the current goal is
mul_SNo_neg_neg x y Hx Hy H1 Hyneg .
We prove the intermediate
claim L1 :
x = 0 .
An exact proof term for the current goal is Hxnp .
An exact proof term for the current goal is H1 .
rewrite the current goal using L1 (from left to right).
rewrite the current goal using
mul_SNo_zeroL y Hy (from left to right).
∎
Proof: Let x, y and z be given.
Assume Hx Hxnp Hy Hz Hzy .
We will
prove x * y < x * z .
An
exact proof term for the current goal is
neg_mul_SNo_Lt x y z Hx H1 Hy Hz H2 .
We prove the intermediate
claim L1 :
y = z .
An exact proof term for the current goal is H2 .
An exact proof term for the current goal is Hzy .
rewrite the current goal using L1 (from left to right).
We prove the intermediate
claim L1 :
x = 0 .
An exact proof term for the current goal is Hxnp .
An exact proof term for the current goal is H1 .
rewrite the current goal using L1 (from left to right).
We will
prove 0 * y ≤ 0 * z .
rewrite the current goal using
mul_SNo_zeroL y Hy (from left to right).
rewrite the current goal using
mul_SNo_zeroL z Hz (from left to right).
∎
Proof: Let x be given.
Assume Hx .
rewrite the current goal using
mul_SNo_zeroR x Hx (from right to left).
We will
prove x * 0 < x * x .
rewrite the current goal using
mul_SNo_zeroR x Hx (from right to left).
We will
prove x * 0 ≤ x * x .
∎
Proof: Let x be given.
Assume Hx .
Apply orIR to the current goal.
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is H1 .
An exact proof term for the current goal is H1 .
Apply orIL to the current goal.
An exact proof term for the current goal is H1 .
Apply orIR to the current goal.
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is Hx .
An exact proof term for the current goal is H1 .
An exact proof term for the current goal is H1 .
∎
Proof: Let x and y be given.
Assume Hx Hy Hxpos Hypos .
We will
prove x * x < x * x .
rewrite the current goal using H1 (from left to right) at position 2.
We will
prove x * x < y * y .
An
exact proof term for the current goal is
pos_mul_SNo_Lt2 x x y y Hx Hx Hy Hy Hxpos Hxpos H2 H2 .
An exact proof term for the current goal is H2 .
We will
prove x * x < x * x .
rewrite the current goal using H1 (from left to right) at position 1.
We will
prove y * y < x * x .
An
exact proof term for the current goal is
pos_mul_SNo_Lt2 y y x x Hy Hy Hx Hx Hypos Hypos H2 H2 .
∎
Proof: Let x and y be given.
Assume Hx Hy Hxnn Hynn .
rewrite the current goal using H2 (from right to left).
An exact proof term for the current goal is H4 .
rewrite the current goal using H3 (from right to left) at position 2.
An exact proof term for the current goal is H4 .
rewrite the current goal using H1 (from right to left).
Use symmetry.
An exact proof term for the current goal is H3 .
rewrite the current goal using H2 (from left to right) at position 2.
An exact proof term for the current goal is H3 .
∎
Theorem. (
SNo_foil ) The following is provable:
Proof: Let x, y, z and w be given.
Assume Hx Hy Hz Hw .
Use transitivity with
(x + y ) * z + (x + y ) * w ,
(x * z + y * z ) + (x + y ) * w , and
(x * z + y * z ) + (x * w + y * w ) .
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_distrR x y z Hx Hy Hz .
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_distrR x y w Hx Hy Hw .
We will
prove (x * z + y * z ) + (x * w + y * w ) = x * z + (x * w + (y * z + y * w ) ) .
We will
prove (x * z + x * w ) + (y * z + y * w ) = x * z + (x * w + (y * z + y * w ) ) .
Use symmetry.
∎
Proof: Let x and y be given.
Assume Hx Hy .
We will
prove - (x * (- y ) ) = x * y .
We will
prove - - (x * y ) = x * y .
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz .
rewrite the current goal using
mul_SNo_assoc x y z Hx Hy Hz (from left to right).
rewrite the current goal using
mul_SNo_assoc y x z Hy Hx Hz (from left to right).
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_com x y Hx Hy .
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz .
rewrite the current goal using
mul_SNo_assoc x y z Hx Hy Hz (from right to left).
rewrite the current goal using
mul_SNo_assoc x z y Hx Hz Hy (from right to left).
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_com y z Hy Hz .
∎
Proof: Let x, y, z and w be given.
Assume Hx Hy Hz Hw .
We will
prove ((x * y ) * z ) * w = (x * z ) * (y * w ) .
We will
prove ((x * z ) * y ) * w = (x * z ) * (y * w ) .
Use symmetry.
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz .
We will
prove x * (y * z ) = z * (x * y ) .
Use transitivity with
x * (z * y ) ,
(x * z ) * y , and
(z * x ) * y .
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_com y z Hy Hz .
An
exact proof term for the current goal is
mul_SNo_assoc x z y Hx Hz Hy .
Use f_equal.
An
exact proof term for the current goal is
mul_SNo_com x z Hx Hz .
Use symmetry.
An
exact proof term for the current goal is
mul_SNo_assoc z x y Hz Hx Hy .
∎
Proof: Let x, y, z and w be given.
Assume Hx Hy Hz Hw .
We will
prove x * w * y * z = w * x * y * z .
∎
Proof: Let x, y, z and w be given.
Assume Hx Hy Hz Hw .
We prove the intermediate
claim Lmy :
SNo (- y ) .
We prove the intermediate
claim Lmw :
SNo (- w ) .
rewrite the current goal using
SNo_foil x (- y ) z (- w ) Hx Lmy Hz Lmw (from left to right).
Use reflexivity.
∎
Proof: Let x, y and z be given.
Assume Hx Hxn0 Hy Hz Hxyz .
We will
prove x * y < x * y .
rewrite the current goal using Hxyz (from left to right) at position 1.
We will
prove x * z < x * y .
An
exact proof term for the current goal is
neg_mul_SNo_Lt x z y Hx H2 Hz Hy H1 .
An exact proof term for the current goal is Hxn0 H2 .
We will
prove x * y < x * y .
rewrite the current goal using Hxyz (from left to right) at position 2.
We will
prove x * y < x * z .
An
exact proof term for the current goal is
pos_mul_SNo_Lt x y z Hx H2 Hy Hz H1 .
Assume H1 .
An exact proof term for the current goal is H1 .
We will
prove x * y < x * y .
rewrite the current goal using Hxyz (from left to right) at position 2.
We will
prove x * y < x * z .
An
exact proof term for the current goal is
neg_mul_SNo_Lt x y z Hx H2 Hy Hz H1 .
An exact proof term for the current goal is Hxn0 H2 .
We will
prove x * y < x * y .
rewrite the current goal using Hxyz (from left to right) at position 1.
We will
prove x * z < x * y .
An
exact proof term for the current goal is
pos_mul_SNo_Lt x z y Hx H2 Hz Hy H1 .
∎
Theorem. (
mul_SNoCutP_lem ) The following is provable:
∀Lx Rx Ly Ry x y, SNoCutP Lx Rx → SNoCutP Ly Ry → x = SNoCut Lx Rx → y = SNoCut Ly Ry → SNoCutP ({ w 0 * y + x * w 1 + - w 0 * w 1 | w ∈ Lx ⨯ Ly } ∪ { z 0 * y + x * z 1 + - z 0 * z 1 | z ∈ Rx ⨯ Ry } ) ({ w 0 * y + x * w 1 + - w 0 * w 1 | w ∈ Lx ⨯ Ry } ∪ { z 0 * y + x * z 1 + - z 0 * z 1 | z ∈ Rx ⨯ Ly } ) ∧ x * y = SNoCut ({ w 0 * y + x * w 1 + - w 0 * w 1 | w ∈ Lx ⨯ Ly } ∪ { z 0 * y + x * z 1 + - z 0 * z 1 | z ∈ Rx ⨯ Ry } ) ({ w 0 * y + x * w 1 + - w 0 * w 1 | w ∈ Lx ⨯ Ry } ∪ { z 0 * y + x * z 1 + - z 0 * z 1 | z ∈ Rx ⨯ Ly } ) ∧ ∀q : prop , (∀LxLy' RxRy' LxRy' RxLy', (∀ u ∈ LxLy' , ∀p : prop , (∀ w ∈ Lx , ∀ w' ∈ Ly , SNo w → SNo w' → w < x → w' < y → u = w * y + x * w' + - w * w' → p ) → p ) → (∀ w ∈ Lx , ∀ w' ∈ Ly , w * y + x * w' + - w * w' ∈ LxLy' ) → (∀ u ∈ RxRy' , ∀p : prop , (∀ z ∈ Rx , ∀ z' ∈ Ry , SNo z → SNo z' → x < z → y < z' → u = z * y + x * z' + - z * z' → p ) → p ) → (∀ z ∈ Rx , ∀ z' ∈ Ry , z * y + x * z' + - z * z' ∈ RxRy' ) → (∀ u ∈ LxRy' , ∀p : prop , (∀ w ∈ Lx , ∀ z ∈ Ry , SNo w → SNo z → w < x → y < z → u = w * y + x * z + - w * z → p ) → p ) → (∀ w ∈ Lx , ∀ z ∈ Ry , w * y + x * z + - w * z ∈ LxRy' ) → (∀ u ∈ RxLy' , ∀p : prop , (∀ z ∈ Rx , ∀ w ∈ Ly , SNo z → SNo w → x < z → w < y → u = z * y + x * w + - z * w → p ) → p ) → (∀ z ∈ Rx , ∀ w ∈ Ly , z * y + x * w + - z * w ∈ RxLy' ) → SNoCutP (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) → x * y = SNoCut (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) → q ) → q
Proof: Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye .
Apply HLRx to the current goal.
Assume H .
Apply H to the current goal.
Apply HLRy to the current goal.
Assume H .
Apply H to the current goal.
rewrite the current goal using Hxe (from right to left).
rewrite the current goal using Hye (from right to left).
We prove the intermediate
claim LxLy'E :
∀ u ∈ LxLy' , ∀p : prop , (∀ w ∈ Lx , ∀ w' ∈ Ly , SNo w → SNo w' → w < x → w' < y → u = w * y + x * w' + - w * w' → p ) → p .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp .
Let ww' be given.
We prove the intermediate
claim Lww'0 :
ww' 0 ∈ Lx .
An
exact proof term for the current goal is
ap0_Sigma Lx (λ_ ⇒ Ly ) ww' Hww' .
We prove the intermediate
claim Lww'1 :
ww' 1 ∈ Ly .
An
exact proof term for the current goal is
ap1_Sigma Lx (λ_ ⇒ Ly ) ww' Hww' .
Apply Hp (ww' 0 ) Lww'0 (ww' 1 ) Lww'1 to the current goal.
We will
prove SNo (ww' 0 ) .
An
exact proof term for the current goal is
HLRx1 (ww' 0 ) Lww'0 .
We will
prove SNo (ww' 1 ) .
An
exact proof term for the current goal is
HLRy1 (ww' 1 ) Lww'1 .
An
exact proof term for the current goal is
Hx3 (ww' 0 ) Lww'0 .
An
exact proof term for the current goal is
Hy3 (ww' 1 ) Lww'1 .
An exact proof term for the current goal is Hue .
We prove the intermediate
claim LxLy'I :
∀ w ∈ Lx , ∀ w' ∈ Ly , w * y + x * w' + - w * w' ∈ LxLy' .
Let w be given.
Assume Hw .
Let w' be given.
Assume Hw' .
rewrite the current goal using
tuple_2_0_eq w w' (from right to left).
rewrite the current goal using
tuple_2_1_eq w w' (from right to left) at positions 2 and 4.
Apply ReplI (Lx ⨯ Ly ) (λw ⇒ w 0 * y + x * w 1 + - w 0 * w 1 ) (w ,w' ) to the current goal.
We will
prove (w ,w' ) ∈ Lx ⨯ Ly .
An
exact proof term for the current goal is
tuple_2_setprod Lx Ly w Hw w' Hw' .
We prove the intermediate
claim RxRy'E :
∀ u ∈ RxRy' , ∀p : prop , (∀ z ∈ Rx , ∀ z' ∈ Ry , SNo z → SNo z' → x < z → y < z' → u = z * y + x * z' + - z * z' → p ) → p .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp .
Let zz' be given.
We prove the intermediate
claim Lzz'0 :
zz' 0 ∈ Rx .
An
exact proof term for the current goal is
ap0_Sigma Rx (λ_ ⇒ Ry ) zz' Hzz' .
We prove the intermediate
claim Lzz'1 :
zz' 1 ∈ Ry .
An
exact proof term for the current goal is
ap1_Sigma Rx (λ_ ⇒ Ry ) zz' Hzz' .
Apply Hp (zz' 0 ) Lzz'0 (zz' 1 ) Lzz'1 to the current goal.
We will
prove SNo (zz' 0 ) .
An
exact proof term for the current goal is
HLRx2 (zz' 0 ) Lzz'0 .
We will
prove SNo (zz' 1 ) .
An
exact proof term for the current goal is
HLRy2 (zz' 1 ) Lzz'1 .
An
exact proof term for the current goal is
Hx4 (zz' 0 ) Lzz'0 .
An
exact proof term for the current goal is
Hy4 (zz' 1 ) Lzz'1 .
An exact proof term for the current goal is Hue .
We prove the intermediate
claim RxRy'I :
∀ z ∈ Rx , ∀ z' ∈ Ry , z * y + x * z' + - z * z' ∈ RxRy' .
Let z be given.
Assume Hz .
Let z' be given.
Assume Hz' .
rewrite the current goal using
tuple_2_0_eq z z' (from right to left).
rewrite the current goal using
tuple_2_1_eq z z' (from right to left) at positions 2 and 4.
Apply ReplI (Rx ⨯ Ry ) (λz ⇒ z 0 * y + x * z 1 + - z 0 * z 1 ) (z ,z' ) to the current goal.
We will
prove (z ,z' ) ∈ Rx ⨯ Ry .
An
exact proof term for the current goal is
tuple_2_setprod Rx Ry z Hz z' Hz' .
We prove the intermediate
claim LxRy'E :
∀ u ∈ LxRy' , ∀p : prop , (∀ w ∈ Lx , ∀ z ∈ Ry , SNo w → SNo z → w < x → y < z → u = w * y + x * z + - w * z → p ) → p .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp .
Let wz be given.
We prove the intermediate
claim Lwz0 :
wz 0 ∈ Lx .
An
exact proof term for the current goal is
ap0_Sigma Lx (λ_ ⇒ Ry ) wz Hwz .
We prove the intermediate
claim Lwz1 :
wz 1 ∈ Ry .
An
exact proof term for the current goal is
ap1_Sigma Lx (λ_ ⇒ Ry ) wz Hwz .
Apply Hp (wz 0 ) Lwz0 (wz 1 ) Lwz1 to the current goal.
We will
prove SNo (wz 0 ) .
An
exact proof term for the current goal is
HLRx1 (wz 0 ) Lwz0 .
We will
prove SNo (wz 1 ) .
An
exact proof term for the current goal is
HLRy2 (wz 1 ) Lwz1 .
An
exact proof term for the current goal is
Hx3 (wz 0 ) Lwz0 .
An
exact proof term for the current goal is
Hy4 (wz 1 ) Lwz1 .
An exact proof term for the current goal is Hue .
We prove the intermediate
claim LxRy'I :
∀ w ∈ Lx , ∀ z ∈ Ry , w * y + x * z + - w * z ∈ LxRy' .
Let w be given.
Assume Hw .
Let z be given.
Assume Hz .
rewrite the current goal using
tuple_2_0_eq w z (from right to left).
rewrite the current goal using
tuple_2_1_eq w z (from right to left) at positions 2 and 4.
Apply ReplI (Lx ⨯ Ry ) (λw ⇒ w 0 * y + x * w 1 + - w 0 * w 1 ) (w ,z ) to the current goal.
We will
prove (w ,z ) ∈ Lx ⨯ Ry .
An
exact proof term for the current goal is
tuple_2_setprod Lx Ry w Hw z Hz .
We prove the intermediate
claim RxLy'E :
∀ u ∈ RxLy' , ∀p : prop , (∀ z ∈ Rx , ∀ w ∈ Ly , SNo z → SNo w → x < z → w < y → u = z * y + x * w + - z * w → p ) → p .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp .
Let zw be given.
We prove the intermediate
claim Lzw0 :
zw 0 ∈ Rx .
An
exact proof term for the current goal is
ap0_Sigma Rx (λ_ ⇒ Ly ) zw Hzw .
We prove the intermediate
claim Lzw1 :
zw 1 ∈ Ly .
An
exact proof term for the current goal is
ap1_Sigma Rx (λ_ ⇒ Ly ) zw Hzw .
Apply Hp (zw 0 ) Lzw0 (zw 1 ) Lzw1 to the current goal.
We will
prove SNo (zw 0 ) .
An
exact proof term for the current goal is
HLRx2 (zw 0 ) Lzw0 .
We will
prove SNo (zw 1 ) .
An
exact proof term for the current goal is
HLRy1 (zw 1 ) Lzw1 .
An
exact proof term for the current goal is
Hx4 (zw 0 ) Lzw0 .
An
exact proof term for the current goal is
Hy3 (zw 1 ) Lzw1 .
An exact proof term for the current goal is Hue .
We prove the intermediate
claim RxLy'I :
∀ z ∈ Rx , ∀ w ∈ Ly , z * y + x * w + - z * w ∈ RxLy' .
Let z be given.
Assume Hz .
Let w be given.
Assume Hw .
rewrite the current goal using
tuple_2_0_eq z w (from right to left).
rewrite the current goal using
tuple_2_1_eq z w (from right to left) at positions 2 and 4.
Apply ReplI (Rx ⨯ Ly ) (λw ⇒ w 0 * y + x * w 1 + - w 0 * w 1 ) (z ,w ) to the current goal.
We will
prove (z ,w ) ∈ Rx ⨯ Ly .
An
exact proof term for the current goal is
tuple_2_setprod Rx Ly z Hz w Hw .
We prove the intermediate
claim L1 :
SNoCutP (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) .
Apply and3I to the current goal.
Let w be given.
Apply LxLy'E w Hw to the current goal.
Let w' be given.
Let w'' be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
We will
prove SNo (w' * y + x * w'' + - w' * w'' ) .
An
exact proof term for the current goal is
SNo_mul_SNo w' y Hw'1 Hy1 .
An
exact proof term for the current goal is
SNo_mul_SNo x w'' Hx1 Hw''1 .
An
exact proof term for the current goal is
SNo_mul_SNo w' w'' Hw'1 Hw''1 .
Apply RxRy'E w Hw to the current goal.
Let z' be given.
Let z'' be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
We will
prove SNo (z' * y + x * z'' + - z' * z'' ) .
An
exact proof term for the current goal is
SNo_mul_SNo z' y Hz'1 Hy1 .
An
exact proof term for the current goal is
SNo_mul_SNo x z'' Hx1 Hz''1 .
An
exact proof term for the current goal is
SNo_mul_SNo z' z'' Hz'1 Hz''1 .
Let z be given.
Apply LxRy'E z Hz to the current goal.
Let w' be given.
Let z' be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove SNo (w' * y + x * z' + - w' * z' ) .
An
exact proof term for the current goal is
SNo_mul_SNo w' y Hw'1 Hy1 .
An
exact proof term for the current goal is
SNo_mul_SNo x z' Hx1 Hz'1 .
An
exact proof term for the current goal is
SNo_mul_SNo w' z' Hw'1 Hz'1 .
Apply RxLy'E z Hz to the current goal.
Let z' be given.
Let w' be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove SNo (z' * y + x * w' + - z' * w' ) .
An
exact proof term for the current goal is
SNo_mul_SNo z' y Hz'1 Hy1 .
An
exact proof term for the current goal is
SNo_mul_SNo x w' Hx1 Hw'1 .
An
exact proof term for the current goal is
SNo_mul_SNo z' w' Hz'1 Hw'1 .
Let w be given.
Apply LxLy'E w Hw to the current goal.
Let w' be given.
Let w'' be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
Let z be given.
Apply LxRy'E z Hz to the current goal.
Let w''' be given.
Let z' be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove w' * y + x * w'' + - w' * w'' < w''' * y + x * z' + - w''' * z' .
We will
prove w' * y + x * w'' + w''' * z' < w''' * y + x * z' + w' * w'' .
We will
prove w' * y + x * w'' + w''' * z' < x * y + w' * w'' + w''' * z' .
We will
prove w''' * z' + w' * y + x * w'' < w''' * z' + x * y + w' * w'' .
We will
prove w' * y + x * w'' < x * y + w' * w'' .
An
exact proof term for the current goal is
mul_SNo_Lt x y w' w'' Hx1 Hy1 Hw'1 Hw''1 Hw'2 Hw''2 .
We will
prove x * y + w' * w'' + w''' * z' < w''' * y + x * z' + w' * w'' .
We will
prove w' * w'' + x * y + w''' * z' < w' * w'' + w''' * y + x * z' .
We will
prove x * y + w''' * z' < w''' * y + x * z' .
We will
prove w''' * z' + x * y < x * z' + w''' * y .
An
exact proof term for the current goal is
mul_SNo_Lt x z' w''' y Hx1 Hz'1 Hw'''1 Hy1 Hw'''2 Hz'2 .
Apply RxLy'E z Hz to the current goal.
Let z' be given.
Let w''' be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove w' * y + x * w'' + - w' * w'' < z' * y + x * w''' + - z' * w''' .
We will
prove w' * y + x * w'' + z' * w''' < z' * y + x * w''' + w' * w'' .
We will
prove w' * y + x * w'' + z' * w''' < z' * w''' + x * y + w' * w'' .
We will
prove z' * w''' + w' * y + x * w'' < z' * w''' + x * y + w' * w'' .
We will
prove w' * y + x * w'' < x * y + w' * w'' .
An
exact proof term for the current goal is
mul_SNo_Lt x y w' w'' Hx1 Hy1 Hw'1 Hw''1 Hw'2 Hw''2 .
We will
prove z' * w''' + x * y + w' * w'' < z' * y + x * w''' + w' * w'' .
We will
prove w' * w'' + z' * w''' + x * y < w' * w'' + z' * y + x * w''' .
We will
prove z' * w''' + x * y < z' * y + x * w''' .
We will
prove x * y + z' * w''' < z' * y + x * w''' .
An
exact proof term for the current goal is
mul_SNo_Lt z' y x w''' Hz'1 Hy1 Hx1 Hw'''1 Hz'2 Hw'''2 .
Apply RxRy'E w Hw to the current goal.
Let z' be given.
Let z'' be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
Let z be given.
Apply LxRy'E z Hz to the current goal.
Let w' be given.
Let z''' be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove z' * y + x * z'' + - z' * z'' < w' * y + x * z''' + - w' * z''' .
We will
prove z' * y + x * z'' + w' * z''' < w' * y + x * z''' + z' * z'' .
We will
prove z' * y + x * z'' + w' * z''' < w' * z''' + z' * z'' + x * y .
We will
prove w' * z''' + z' * y + x * z'' < w' * z''' + z' * z'' + x * y .
We will
prove z' * y + x * z'' < z' * z'' + x * y .
An
exact proof term for the current goal is
mul_SNo_Lt z' z'' x y Hz'1 Hz''1 Hx1 Hy1 Hz'2 Hz''2 .
We will
prove w' * z''' + z' * z'' + x * y < w' * y + x * z''' + z' * z'' .
We will
prove z' * z'' + w' * z''' + x * y < z' * z'' + w' * y + x * z''' .
We will
prove w' * z''' + x * y < w' * y + x * z''' .
An
exact proof term for the current goal is
mul_SNo_Lt x z''' w' y Hx1 Hz'''1 Hw'1 Hy1 Hw'2 Hz'''2 .
Apply RxLy'E z Hz to the current goal.
Let z''' be given.
Let w' be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove z' * y + x * z'' + - z' * z'' < z''' * y + x * w' + - z''' * w' .
We will
prove z' * y + x * z'' + z''' * w' < z''' * y + x * w' + z' * z'' .
We will
prove z' * y + x * z'' + z''' * w' < z''' * w' + z' * z'' + x * y .
We will
prove z''' * w' + x * z'' + z' * y < z''' * w' + z' * z'' + x * y .
We will
prove x * z'' + z' * y < z' * z'' + x * y .
An
exact proof term for the current goal is
mul_SNo_Lt z' z'' x y Hz'1 Hz''1 Hx1 Hy1 Hz'2 Hz''2 .
We will
prove z''' * w' + z' * z'' + x * y < z''' * y + x * w' + z' * z'' .
We will
prove z' * z'' + z''' * w' + x * y < z' * z'' + z''' * y + x * w' .
We will
prove z''' * w' + x * y < z''' * y + x * w' .
An
exact proof term for the current goal is
mul_SNo_Lt z''' y x w' Hz'''1 Hy1 Hx1 Hw'1 Hz'''2 Hw'2 .
We prove the intermediate
claim Lxyeq :
x * y = SNoCut (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) .
Set v to be the term
SNoCut (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) .
Let L' and R' be given.
rewrite the current goal using HLR'eq (from left to right).
Apply SNoCut_ext L' R' (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) to the current goal.
An exact proof term for the current goal is HLR' .
An exact proof term for the current goal is L1 .
We will
prove ∀ w ∈ L' , w < v .
Let w be given.
Assume Hw .
Apply HL'E w Hw to the current goal.
Let w0 be given.
Let w1 be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
We will
prove w0 * y + x * w1 + - w0 * w1 < v .
Apply SNoL_E x Hx1 w0 Hw0 to the current goal.
Assume Hw0a Hw0b Hw0c .
Apply SNoL_E y Hy1 w1 Hw1 to the current goal.
Assume Hw1a Hw1b Hw1c .
We prove the intermediate
claim L2 :
∃ w0' ∈ Lx , w0 ≤ w0' .
Apply dneg to the current goal.
Assume HC .
Apply Hx5 w0 Hw0a to the current goal.
We will
prove ∀ w' ∈ Lx , w' < w0 .
Let w' be given.
Assume Hw' .
Apply SNoLtLe_or w' w0 (HLRx1 w' Hw' ) Hw0a to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw' .
An exact proof term for the current goal is H .
We will
prove ∀ z' ∈ Rx , w0 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLt_tra w0 x z' Hw0a Hx1 (HLRx2 z' Hz' ) to the current goal.
An exact proof term for the current goal is Hw0c .
An exact proof term for the current goal is Hx4 z' Hz' .
Apply L2a to the current goal.
Assume _ .
Apply Hxw0 to the current goal.
An exact proof term for the current goal is Hw0b .
We prove the intermediate
claim L3 :
∃ w1' ∈ Ly , w1 ≤ w1' .
Apply dneg to the current goal.
Assume HC .
Apply Hy5 w1 Hw1a to the current goal.
We will
prove ∀ w' ∈ Ly , w' < w1 .
Let w' be given.
Assume Hw' .
Apply SNoLtLe_or w' w1 (HLRy1 w' Hw' ) Hw1a to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw' .
An exact proof term for the current goal is H .
We will
prove ∀ z' ∈ Ry , w1 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLt_tra w1 y z' Hw1a Hy1 (HLRy2 z' Hz' ) to the current goal.
An exact proof term for the current goal is Hw1c .
An exact proof term for the current goal is Hy4 z' Hz' .
Apply L3a to the current goal.
Assume _ .
Apply Hyw1 to the current goal.
An exact proof term for the current goal is Hw1b .
Apply L2 to the current goal.
Let w0' be given.
Assume H .
Apply H to the current goal.
Apply L3 to the current goal.
Let w1' be given.
Assume H .
Apply H to the current goal.
We will
prove w0 * y + x * w1 + - w0 * w1 < v .
We will
prove w0 * y + x * w1 + - w0 * w1 ≤ w0' * y + x * w1' + - w0' * w1' .
We will
prove w0 * y + x * w1 + w0' * w1' ≤ w0' * y + x * w1' + w0 * w1 .
We will
prove w0 * y + x * w1 + w0' * w1' ≤ w0 * y + x * w1' + w0' * w1 .
We will
prove x * w1 + w0' * w1' ≤ x * w1' + w0' * w1 .
We will
prove w0' * w1' + x * w1 ≤ x * w1' + w0' * w1 .
Apply mul_SNo_Le x w1' w0' w1 Hx1 (HLRy1 w1' Hw1'1 ) (HLRx1 w0' Hw0'1 ) Hw1a to the current goal.
An exact proof term for the current goal is Hx3 w0' Hw0'1 .
An exact proof term for the current goal is Hw1'2 .
We will
prove w0 * y + x * w1' + w0' * w1 ≤ w0' * y + x * w1' + w0 * w1 .
We will
prove w0 * y + w0' * w1 ≤ w0' * y + w0 * w1 .
Apply mul_SNo_Le w0' y w0 w1 (HLRx1 w0' Hw0'1 ) Hy1 Hw0a Hw1a to the current goal.
An exact proof term for the current goal is Hw0'2 .
An exact proof term for the current goal is Hw1c .
We will
prove w0' * y + x * w1' + - w0' * w1' < v .
Apply Hv3 to the current goal.
We will
prove w0' * y + x * w1' + - w0' * w1' ∈ LxLy' ∪ RxRy' .
Apply LxLy'I to the current goal.
An exact proof term for the current goal is Hw0'1 .
An exact proof term for the current goal is Hw1'1 .
Let z0 be given.
Let z1 be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
We will
prove z0 * y + x * z1 + - z0 * z1 < v .
Apply SNoR_E x Hx1 z0 Hz0 to the current goal.
Assume Hz0a Hz0b Hz0c .
Apply SNoR_E y Hy1 z1 Hz1 to the current goal.
Assume Hz1a Hz1b Hz1c .
We prove the intermediate
claim L4 :
∃ z0' ∈ Rx , z0' ≤ z0 .
Apply dneg to the current goal.
Assume HC .
Apply Hx5 z0 Hz0a to the current goal.
We will
prove ∀ w' ∈ Lx , w' < z0 .
Let w' be given.
Assume Hw' .
Apply SNoLt_tra w' x z0 (HLRx1 w' Hw' ) Hx1 Hz0a to the current goal.
An exact proof term for the current goal is Hx3 w' Hw' .
An exact proof term for the current goal is Hz0c .
We will
prove ∀ z' ∈ Rx , z0 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLtLe_or z0 z' Hz0a (HLRx2 z' Hz' ) to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz' .
An exact proof term for the current goal is H .
Apply L4a to the current goal.
Assume _ .
Apply Hxz0 to the current goal.
An exact proof term for the current goal is Hz0b .
We prove the intermediate
claim L5 :
∃ z1' ∈ Ry , z1' ≤ z1 .
Apply dneg to the current goal.
Assume HC .
Apply Hy5 z1 Hz1a to the current goal.
We will
prove ∀ w' ∈ Ly , w' < z1 .
Let w' be given.
Assume Hw' .
Apply SNoLt_tra w' y z1 (HLRy1 w' Hw' ) Hy1 Hz1a to the current goal.
An exact proof term for the current goal is Hy3 w' Hw' .
An exact proof term for the current goal is Hz1c .
We will
prove ∀ z' ∈ Ry , z1 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLtLe_or z1 z' Hz1a (HLRy2 z' Hz' ) to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz' .
An exact proof term for the current goal is H .
Apply L5a to the current goal.
Assume _ .
Apply Hyz1 to the current goal.
An exact proof term for the current goal is Hz1b .
Apply L4 to the current goal.
Let z0' be given.
Assume H .
Apply H to the current goal.
Apply L5 to the current goal.
Let z1' be given.
Assume H .
Apply H to the current goal.
We will
prove z0 * y + x * z1 + - z0 * z1 < v .
We will
prove z0 * y + x * z1 + - z0 * z1 ≤ z0' * y + x * z1' + - z0' * z1' .
We will
prove z0 * y + x * z1 + z0' * z1' ≤ z0' * y + x * z1' + z0 * z1 .
We will
prove z0 * y + x * z1 + z0' * z1' ≤ z0 * y + z0' * z1 + x * z1' .
We will
prove x * z1 + z0' * z1' ≤ z0' * z1 + x * z1' .
Apply mul_SNo_Le z0' z1 x z1' (HLRx2 z0' Hz0'1 ) Hz1a Hx1 (HLRy2 z1' Hz1'1 ) to the current goal.
An exact proof term for the current goal is Hx4 z0' Hz0'1 .
An exact proof term for the current goal is Hz1'2 .
We will
prove z0 * y + z0' * z1 + x * z1' ≤ z0' * y + x * z1' + z0 * z1 .
We will
prove x * z1' + z0 * y + z0' * z1 ≤ x * z1' + z0' * y + z0 * z1 .
We will
prove z0 * y + z0' * z1 ≤ z0' * y + z0 * z1 .
Apply mul_SNo_Le z0 z1 z0' y Hz0a Hz1a (HLRx2 z0' Hz0'1 ) Hy1 to the current goal.
An exact proof term for the current goal is Hz0'2 .
An exact proof term for the current goal is Hz1c .
We will
prove z0' * y + x * z1' + - z0' * z1' < v .
Apply Hv3 to the current goal.
We will
prove z0' * y + x * z1' + - z0' * z1' ∈ LxLy' ∪ RxRy' .
Apply RxRy'I to the current goal.
An exact proof term for the current goal is Hz0'1 .
An exact proof term for the current goal is Hz1'1 .
We will
prove ∀ z ∈ R' , v < z .
Let z be given.
Assume Hz .
Apply HR'E z Hz to the current goal.
Let w0 be given.
Let z1 be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove v < w0 * y + x * z1 + - w0 * z1 .
Apply SNoL_E x Hx1 w0 Hw0 to the current goal.
Assume Hw0a Hw0b Hw0c .
Apply SNoR_E y Hy1 z1 Hz1 to the current goal.
Assume Hz1a Hz1b Hz1c .
We prove the intermediate
claim L6 :
∃ w0' ∈ Lx , w0 ≤ w0' .
Apply dneg to the current goal.
Assume HC .
Apply Hx5 w0 Hw0a to the current goal.
We will
prove ∀ w' ∈ Lx , w' < w0 .
Let w' be given.
Assume Hw' .
Apply SNoLtLe_or w' w0 (HLRx1 w' Hw' ) Hw0a to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw' .
An exact proof term for the current goal is H .
We will
prove ∀ z' ∈ Rx , w0 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLt_tra w0 x z' Hw0a Hx1 (HLRx2 z' Hz' ) to the current goal.
An exact proof term for the current goal is Hw0c .
An exact proof term for the current goal is Hx4 z' Hz' .
Apply L6a to the current goal.
Assume _ .
Apply Hxw0 to the current goal.
An exact proof term for the current goal is Hw0b .
We prove the intermediate
claim L7 :
∃ z1' ∈ Ry , z1' ≤ z1 .
Apply dneg to the current goal.
Assume HC .
Apply Hy5 z1 Hz1a to the current goal.
We will
prove ∀ w' ∈ Ly , w' < z1 .
Let w' be given.
Assume Hw' .
Apply SNoLt_tra w' y z1 (HLRy1 w' Hw' ) Hy1 Hz1a to the current goal.
An exact proof term for the current goal is Hy3 w' Hw' .
An exact proof term for the current goal is Hz1c .
We will
prove ∀ z' ∈ Ry , z1 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLtLe_or z1 z' Hz1a (HLRy2 z' Hz' ) to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz' .
An exact proof term for the current goal is H .
Apply L7a to the current goal.
Assume _ .
Apply Hyz1 to the current goal.
An exact proof term for the current goal is Hz1b .
Apply L6 to the current goal.
Let w0' be given.
Assume H .
Apply H to the current goal.
Apply L7 to the current goal.
Let z1' be given.
Assume H .
Apply H to the current goal.
We will
prove v < w0 * y + x * z1 + - w0 * z1 .
We will
prove v < w0' * y + x * z1' + - w0' * z1' .
Apply Hv4 to the current goal.
We will
prove w0' * y + x * z1' + - w0' * z1' ∈ LxRy' ∪ RxLy' .
Apply LxRy'I to the current goal.
An exact proof term for the current goal is Hw0'1 .
An exact proof term for the current goal is Hz1'1 .
We will
prove w0' * y + x * z1' + - w0' * z1' ≤ w0 * y + x * z1 + - w0 * z1 .
We will
prove w0' * y + x * z1' + w0 * z1 ≤ w0 * y + x * z1 + w0' * z1' .
We will
prove w0' * y + x * z1' + w0 * z1 ≤ x * z1' + w0' * z1 + w0 * y .
We will
prove w0' * y + w0 * z1 ≤ w0' * z1 + w0 * y .
Apply mul_SNo_Le w0' z1 w0 y (HLRx1 w0' Hw0'1 ) Hz1a Hw0a Hy1 to the current goal.
An exact proof term for the current goal is Hw0'2 .
An exact proof term for the current goal is Hz1c .
We will
prove x * z1' + w0' * z1 + w0 * y ≤ w0 * y + x * z1 + w0' * z1' .
We will
prove w0 * y + x * z1' + w0' * z1 ≤ w0 * y + x * z1 + w0' * z1' .
We will
prove x * z1' + w0' * z1 ≤ x * z1 + w0' * z1' .
We will
prove w0' * z1 + x * z1' ≤ x * z1 + w0' * z1' .
Apply mul_SNo_Le x z1 w0' z1' Hx1 Hz1a (HLRx1 w0' Hw0'1 ) (HLRy2 z1' Hz1'1 ) to the current goal.
An exact proof term for the current goal is Hx3 w0' Hw0'1 .
An exact proof term for the current goal is Hz1'2 .
Let z0 be given.
Let w1 be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove v < z0 * y + x * w1 + - z0 * w1 .
Apply SNoR_E x Hx1 z0 Hz0 to the current goal.
Assume Hz0a Hz0b Hz0c .
Apply SNoL_E y Hy1 w1 Hw1 to the current goal.
Assume Hw1a Hw1b Hw1c .
We prove the intermediate
claim L8 :
∃ z0' ∈ Rx , z0' ≤ z0 .
Apply dneg to the current goal.
Assume HC .
Apply Hx5 z0 Hz0a to the current goal.
We will
prove ∀ w' ∈ Lx , w' < z0 .
Let w' be given.
Assume Hw' .
Apply SNoLt_tra w' x z0 (HLRx1 w' Hw' ) Hx1 Hz0a to the current goal.
An exact proof term for the current goal is Hx3 w' Hw' .
An exact proof term for the current goal is Hz0c .
We will
prove ∀ z' ∈ Rx , z0 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLtLe_or z0 z' Hz0a (HLRx2 z' Hz' ) to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use z' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz' .
An exact proof term for the current goal is H .
Apply L8a to the current goal.
Assume _ .
Apply Hxz0 to the current goal.
An exact proof term for the current goal is Hz0b .
We prove the intermediate
claim L9 :
∃ w1' ∈ Ly , w1 ≤ w1' .
Apply dneg to the current goal.
Assume HC .
Apply Hy5 w1 Hw1a to the current goal.
We will
prove ∀ w' ∈ Ly , w' < w1 .
Let w' be given.
Assume Hw' .
Apply SNoLtLe_or w' w1 (HLRy1 w' Hw' ) Hw1a to the current goal.
Assume H .
An exact proof term for the current goal is H .
Apply HC to the current goal.
We use w' to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw' .
An exact proof term for the current goal is H .
We will
prove ∀ z' ∈ Ry , w1 < z' .
Let z' be given.
Assume Hz' .
Apply SNoLt_tra w1 y z' Hw1a Hy1 (HLRy2 z' Hz' ) to the current goal.
An exact proof term for the current goal is Hw1c .
An exact proof term for the current goal is Hy4 z' Hz' .
Apply L9a to the current goal.
Assume _ .
Apply Hyw1 to the current goal.
An exact proof term for the current goal is Hw1b .
Apply L8 to the current goal.
Let z0' be given.
Assume H .
Apply H to the current goal.
Apply L9 to the current goal.
Let w1' be given.
Assume H .
Apply H to the current goal.
We will
prove v < z0 * y + x * w1 + - z0 * w1 .
We will
prove v < z0' * y + x * w1' + - z0' * w1' .
Apply Hv4 to the current goal.
We will
prove z0' * y + x * w1' + - z0' * w1' ∈ LxRy' ∪ RxLy' .
Apply RxLy'I to the current goal.
An exact proof term for the current goal is Hz0'1 .
An exact proof term for the current goal is Hw1'1 .
We will
prove z0' * y + x * w1' + - z0' * w1' ≤ z0 * y + x * w1 + - z0 * w1 .
We will
prove z0' * y + x * w1' + z0 * w1 ≤ z0 * y + x * w1 + z0' * w1' .
We will
prove z0' * y + x * w1' + z0 * w1 ≤ z0' * y + x * w1 + z0 * w1' .
We will
prove x * w1' + z0 * w1 ≤ x * w1 + z0 * w1' .
Apply mul_SNo_Le z0 w1' x w1 Hz0a (HLRy1 w1' Hw1'1 ) Hx1 Hw1a to the current goal.
An exact proof term for the current goal is Hz0c .
An exact proof term for the current goal is Hw1'2 .
We will
prove z0' * y + x * w1 + z0 * w1' ≤ z0 * y + x * w1 + z0' * w1' .
We will
prove x * w1 + z0' * y + z0 * w1' ≤ x * w1 + z0 * y + z0' * w1' .
We will
prove z0' * y + z0 * w1' ≤ z0 * y + z0' * w1' .
Apply mul_SNo_Le z0 y z0' w1' Hz0a Hy1 (HLRx2 z0' Hz0'1 ) (HLRy1 w1' Hw1'1 ) to the current goal.
An exact proof term for the current goal is Hz0'2 .
An exact proof term for the current goal is Hy3 w1' Hw1'1 .
rewrite the current goal using HLR'eq (from right to left).
We will
prove ∀ w ∈ LxLy' ∪ RxRy' , w < x * y .
Let w be given.
Apply LxLy'E w Hw to the current goal.
Let w' be given.
Let w'' be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
We will
prove w' * y + x * w'' + - w' * w'' < x * y .
We will
prove w' * y + x * w'' < x * y + w' * w'' .
An
exact proof term for the current goal is
mul_SNo_Lt x y w' w'' Hx1 Hy1 Hw'1 Hw''1 Hw'2 Hw''2 .
Apply RxRy'E w Hw to the current goal.
Let z be given.
Let z' be given.
Assume Hwe .
rewrite the current goal using Hwe (from left to right).
We will
prove z * y + x * z' + - z * z' < x * y .
We will
prove z * y + x * z' < x * y + z * z' .
An
exact proof term for the current goal is
mul_SNo_Lt z z' x y Hz1 Hz'1 Hx1 Hy1 Hz2 Hz'2 .
rewrite the current goal using HLR'eq (from right to left).
We will
prove ∀ z ∈ LxRy' ∪ RxLy' , x * y < z .
Let z be given.
Apply LxRy'E z Hz to the current goal.
Let w be given.
Let z' be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove x * y < w * y + x * z' + - w * z' .
We will
prove x * y + w * z' < w * y + x * z' .
We will
prove w * z' + x * y < x * z' + w * y .
An
exact proof term for the current goal is
mul_SNo_Lt x z' w y Hx1 Hz'1 Hw1 Hy1 Hw2 Hz'2 .
Apply RxLy'E z Hz to the current goal.
Let z' be given.
Let w be given.
Assume Hze .
rewrite the current goal using Hze (from left to right).
We will
prove x * y < z' * y + x * w + - z' * w .
We will
prove x * y + z' * w < z' * y + x * w .
An
exact proof term for the current goal is
mul_SNo_Lt z' y x w Hz'1 Hy1 Hx1 Hw1 Hz'2 Hw2 .
Apply and3I to the current goal.
An exact proof term for the current goal is L1 .
An exact proof term for the current goal is Lxyeq .
Let q be given.
Assume Hq .
Apply Hq LxLy' RxRy' LxRy' RxLy' to the current goal.
An exact proof term for the current goal is LxLy'E .
An exact proof term for the current goal is LxLy'I .
An exact proof term for the current goal is RxRy'E .
An exact proof term for the current goal is RxRy'I .
An exact proof term for the current goal is LxRy'E .
An exact proof term for the current goal is LxRy'I .
An exact proof term for the current goal is RxLy'E .
An exact proof term for the current goal is RxLy'I .
An exact proof term for the current goal is L1 .
An exact proof term for the current goal is Lxyeq .
∎
Proof: Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye .
Apply mul_SNoCutP_lem Lx Rx Ly Ry x y HLRx HLRy Hxe Hye to the current goal.
Assume H _ .
Apply H to the current goal.
Assume H _ .
An exact proof term for the current goal is H .
∎
Proof: Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye .
Apply mul_SNoCutP_lem Lx Rx Ly Ry x y HLRx HLRy Hxe Hye to the current goal.
Assume H _ .
Apply H to the current goal.
Assume _ H .
An exact proof term for the current goal is H .
∎
Theorem. (
mul_SNoCut_abs ) The following is provable:
∀Lx Rx Ly Ry x y, SNoCutP Lx Rx → SNoCutP Ly Ry → x = SNoCut Lx Rx → y = SNoCut Ly Ry → ∀q : prop , (∀LxLy' RxRy' LxRy' RxLy', (∀ u ∈ LxLy' , ∀p : prop , (∀ w ∈ Lx , ∀ w' ∈ Ly , SNo w → SNo w' → w < x → w' < y → u = w * y + x * w' + - w * w' → p ) → p ) → (∀ w ∈ Lx , ∀ w' ∈ Ly , w * y + x * w' + - w * w' ∈ LxLy' ) → (∀ u ∈ RxRy' , ∀p : prop , (∀ z ∈ Rx , ∀ z' ∈ Ry , SNo z → SNo z' → x < z → y < z' → u = z * y + x * z' + - z * z' → p ) → p ) → (∀ z ∈ Rx , ∀ z' ∈ Ry , z * y + x * z' + - z * z' ∈ RxRy' ) → (∀ u ∈ LxRy' , ∀p : prop , (∀ w ∈ Lx , ∀ z ∈ Ry , SNo w → SNo z → w < x → y < z → u = w * y + x * z + - w * z → p ) → p ) → (∀ w ∈ Lx , ∀ z ∈ Ry , w * y + x * z + - w * z ∈ LxRy' ) → (∀ u ∈ RxLy' , ∀p : prop , (∀ z ∈ Rx , ∀ w ∈ Ly , SNo z → SNo w → x < z → w < y → u = z * y + x * w + - z * w → p ) → p ) → (∀ z ∈ Rx , ∀ w ∈ Ly , z * y + x * w + - z * w ∈ RxLy' ) → SNoCutP (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) → x * y = SNoCut (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) → q ) → q
Proof: Let Lx, Rx, Ly, Ry, x and y be given.
Assume HLRx HLRy Hxe Hye .
Apply mul_SNoCutP_lem Lx Rx Ly Ry x y HLRx HLRy Hxe Hye to the current goal.
Assume _ H .
An exact proof term for the current goal is H .
∎
Proof: Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy .
Let x and y be given.
Assume Hx Hy .
Set P1 to be the term
λu ⇒ ∃ v ∈ Lx , ∃ w ∈ Ly , u + v * w ≤ v * y + x * w of type
set → prop .
Set P2 to be the term
λu ⇒ ∃ v ∈ Rx , ∃ w ∈ Ry , u + v * w ≤ v * y + x * w of type
set → prop .
Set P to be the term
λu ⇒ P1 u ∨ P2 u of type
set → prop .
Apply HLRx to the current goal.
Assume H .
Apply H to the current goal.
Assume HLx HRx HLRx' .
Apply HLRy to the current goal.
Assume H .
Apply H to the current goal.
Assume HLy HRy HLRy' .
rewrite the current goal using Hx (from right to left).
rewrite the current goal using Hy (from right to left).
We prove the intermediate
claim Lxy :
SNo (x * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo x y Hx1 Hy1 .
Let u be given.
Apply dneg to the current goal.
We prove the intermediate
claim L1 :
∀ v ∈ Lx , ∀ w ∈ Ly , v * y + x * w < u + v * w .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HLx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HLy w Hw .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
We prove the intermediate
claim L2 :
∀ v ∈ Rx , ∀ w ∈ Ry , v * y + x * w < u + v * w .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HRx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HRy w Hw .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
Apply SNoLtLe_tra u (x * y ) u Hu1 Lxy Hu1 Hu3 to the current goal.
Apply mul_SNoCut_abs Lx Rx Ly Ry x y HLRx HLRy Hx Hy to the current goal.
Let LxLy', RxRy', LxRy' and RxLy' be given.
Assume LxLy'E LxLy'I RxRy'E RxRy'I LxRy'E LxRy'I RxLy'E RxLy'I .
rewrite the current goal using HE (from left to right).
We will
prove SNoCut (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) ≤ u .
rewrite the current goal using
SNo_eta u Hu1 (from left to right).
Let z be given.
rewrite the current goal using
SNo_eta u Hu1 (from right to left).
Apply LxLy'E z Hz to the current goal.
Let w0 be given.
Assume Hw0 .
Let w1 be given.
Assume Hw1 .
Assume Hw0a Hw1a Hw0x Hw1y Hze .
rewrite the current goal using Hze (from left to right).
We will
prove w0 * y + x * w1 + - w0 * w1 < u .
We will
prove w0 * y + x * w1 < u + w0 * w1 .
Assume H .
An exact proof term for the current goal is H .
Apply HNC to the current goal.
We will
prove P1 u ∨ P2 u .
Apply orIL to the current goal.
We use w0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw0 .
We use w1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw1 .
An exact proof term for the current goal is H .
Apply RxRy'E z Hz to the current goal.
Let z0 be given.
Assume Hz0 .
Let z1 be given.
Assume Hz1 .
Assume Hz0a Hz1a Hz0x Hz1y Hze .
rewrite the current goal using Hze (from left to right).
We will
prove z0 * y + x * z1 + - z0 * z1 < u .
We will
prove z0 * y + x * z1 < u + z0 * z1 .
Assume H .
An exact proof term for the current goal is H .
Apply HNC to the current goal.
We will
prove P1 u ∨ P2 u .
Apply orIR to the current goal.
We use z0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz0 .
We use z1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz1 .
An exact proof term for the current goal is H .
Let z be given.
rewrite the current goal using HE (from right to left).
Apply SNoR_E u Hu1 z Hz to the current goal.
We prove the intermediate claim LPz : P z .
Apply IH z to the current goal.
An exact proof term for the current goal is Hz1 .
An exact proof term for the current goal is Hu1 .
An exact proof term for the current goal is Hz2 .
An exact proof term for the current goal is H1 .
Apply LPz to the current goal.
Assume H2 : P1 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HLx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HLy w Hw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Lv1 Lw1 .
We prove the intermediate
claim L3 :
z + v * w < u + v * w .
We will
prove v * y + x * w < u + v * w .
An exact proof term for the current goal is L1 v Hv w Hw .
We prove the intermediate
claim L4 :
z < u .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L4 .
Assume H2 : P2 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HRx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HRy w Hw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Lv1 Lw1 .
We prove the intermediate
claim L5 :
z + v * w < u + v * w .
We will
prove v * y + x * w < u + v * w .
An exact proof term for the current goal is L2 v Hv w Hw .
We prove the intermediate
claim L6 :
z < u .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 Hz3 L6 .
rewrite the current goal using H1 (from right to left).
An exact proof term for the current goal is Hz2 .
An exact proof term for the current goal is H1 .
Let u be given.
Apply SNoL_E (x * y ) Lxy u Hu to the current goal.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3 .
∎
Proof: Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy .
Let x and y be given.
Assume Hx Hy .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp1 Hp2 .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw .
∎
Proof: Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy .
Let x and y be given.
Assume Hx Hy .
Set P1 to be the term
λu ⇒ ∃ v ∈ Lx , ∃ w ∈ Ry , v * y + x * w ≤ u + v * w of type
set → prop .
Set P2 to be the term
λu ⇒ ∃ v ∈ Rx , ∃ w ∈ Ly , v * y + x * w ≤ u + v * w of type
set → prop .
Set P to be the term
λu ⇒ P1 u ∨ P2 u of type
set → prop .
Apply HLRx to the current goal.
Assume H .
Apply H to the current goal.
Assume HLx HRx HLRx' .
Apply HLRy to the current goal.
Assume H .
Apply H to the current goal.
Assume HLy HRy HLRy' .
rewrite the current goal using Hx (from right to left).
rewrite the current goal using Hy (from right to left).
We prove the intermediate
claim Lxy :
SNo (x * y ) .
An
exact proof term for the current goal is
SNo_mul_SNo x y Hx1 Hy1 .
Let u be given.
Apply dneg to the current goal.
We prove the intermediate
claim L1 :
∀ v ∈ Lx , ∀ w ∈ Ry , u + v * w < v * y + x * w .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HLx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HRy w Hw .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIL to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
We prove the intermediate
claim L2 :
∀ v ∈ Rx , ∀ w ∈ Ly , u + v * w < v * y + x * w .
Let v be given.
Assume Hv .
Let w be given.
Assume Hw .
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HRx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HLy w Hw .
Assume H .
An exact proof term for the current goal is H .
Assume H .
Apply HNC to the current goal.
Apply orIR to the current goal.
We use v to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hv .
We use w to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw .
An exact proof term for the current goal is H .
Apply SNoLtLe_tra (x * y ) u (x * y ) Lxy Hu1 Lxy Hu3 to the current goal.
Apply mul_SNoCut_abs Lx Rx Ly Ry x y HLRx HLRy Hx Hy to the current goal.
Let LxLy', RxRy', LxRy' and RxLy' be given.
Assume LxLy'E LxLy'I RxRy'E RxRy'I LxRy'E LxRy'I RxLy'E RxLy'I .
rewrite the current goal using HE (from left to right).
We will
prove u ≤ SNoCut (LxLy' ∪ RxRy' ) (LxRy' ∪ RxLy' ) .
rewrite the current goal using
SNo_eta u Hu1 (from left to right).
Let z be given.
rewrite the current goal using HE (from right to left).
Apply SNoL_E u Hu1 z Hz to the current goal.
We prove the intermediate claim LPz : P z .
Apply IH z to the current goal.
An exact proof term for the current goal is Hz1 .
An exact proof term for the current goal is Hu1 .
An exact proof term for the current goal is Hz2 .
An exact proof term for the current goal is H1 .
Apply LPz to the current goal.
Assume H2 : P1 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HLx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HRy w Hw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Lv1 Lw1 .
We prove the intermediate
claim L3 :
u + v * w < z + v * w .
We will
prove u + v * w < v * y + x * w .
An exact proof term for the current goal is L1 v Hv w Hw .
An exact proof term for the current goal is Hvw .
We prove the intermediate
claim L4 :
u < z .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 L4 Hz3 .
Assume H2 : P2 z .
Apply H2 to the current goal.
Let v be given.
Assume H2 .
Apply H2 to the current goal.
Assume H2 .
Apply H2 to the current goal.
Let w be given.
Assume H2 .
Apply H2 to the current goal.
We prove the intermediate
claim Lv1 :
SNo v .
An exact proof term for the current goal is HRx v Hv .
We prove the intermediate
claim Lw1 :
SNo w .
An exact proof term for the current goal is HLy w Hw .
We prove the intermediate
claim Lvw :
SNo (v * w ) .
An
exact proof term for the current goal is
SNo_mul_SNo v w Lv1 Lw1 .
We prove the intermediate
claim L5 :
u + v * w < z + v * w .
We will
prove u + v * w < v * y + x * w .
An exact proof term for the current goal is L2 v Hv w Hw .
An exact proof term for the current goal is Hvw .
We prove the intermediate
claim L6 :
u < z .
An
exact proof term for the current goal is
SNoLt_tra u z u Hu1 Hz1 Hu1 L6 Hz3 .
rewrite the current goal using H1 (from left to right).
An exact proof term for the current goal is Hz2 .
An exact proof term for the current goal is H1 .
Let z be given.
rewrite the current goal using
SNo_eta u Hu1 (from right to left).
Apply LxRy'E z Hz to the current goal.
Let w0 be given.
Assume Hw0 .
Let w1 be given.
Assume Hw1 .
Assume Hw0a Hw1a Hw0x Hw1y Hze .
rewrite the current goal using Hze (from left to right).
We will
prove u < w0 * y + x * w1 + - w0 * w1 .
We will
prove u + w0 * w1 < w0 * y + x * w1 .
Assume H .
An exact proof term for the current goal is H .
Apply HNC to the current goal.
We will
prove P1 u ∨ P2 u .
Apply orIL to the current goal.
We use w0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw0 .
We use w1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hw1 .
An exact proof term for the current goal is H .
Apply RxLy'E z Hz to the current goal.
Let z0 be given.
Assume Hz0 .
Let z1 be given.
Assume Hz1 .
Assume Hz0a Hz1a Hz0x Hz1y Hze .
rewrite the current goal using Hze (from left to right).
We will
prove u < z0 * y + x * z1 + - z0 * z1 .
We will
prove u + z0 * z1 < z0 * y + x * z1 .
Assume H .
An exact proof term for the current goal is H .
Apply HNC to the current goal.
We will
prove P1 u ∨ P2 u .
Apply orIR to the current goal.
We use z0 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz0 .
We use z1 to witness the existential quantifier.
Apply andI to the current goal.
An exact proof term for the current goal is Hz1 .
An exact proof term for the current goal is H .
Let u be given.
Apply SNoR_E (x * y ) Lxy u Hu to the current goal.
An exact proof term for the current goal is LI u Hu1 Hu2 Hu3 .
∎
Proof: Let Lx, Rx, Ly and Ry be given.
Assume HLRx HLRy .
Let x and y be given.
Assume Hx Hy .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp1 Hp2 .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp1 v Hv w Hw Hvw .
Assume H1 .
Apply H1 to the current goal.
Let v be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hv .
Assume H1 .
Apply H1 to the current goal.
Let w be given.
Assume H1 .
Apply H1 to the current goal.
Assume Hw Hvw .
An exact proof term for the current goal is Hp2 v Hv w Hw Hvw .
∎
End of Section SurrealMul
Beginning of Section SurrealExp
Notation . We use
- as a prefix operator with priority 358 corresponding to applying term
minus_SNo .
Notation . We use
+ as an infix operator with priority 360 and which associates to the right corresponding to applying term
add_SNo .
Notation . We use
* as an infix operator with priority 355 and which associates to the right corresponding to applying term
mul_SNo .
Notation . We use
^ as an infix operator with priority 342 and which associates to the right corresponding to applying term
exp_SNo_nat .
Proof: Let x be given.
Assume Hx .
An
exact proof term for the current goal is
nat_primrec_0 1 (λ_ r ⇒ x * r ) .
∎
Proof: Let x be given.
Assume Hx .
Let n be given.
Assume Hn .
An
exact proof term for the current goal is
nat_primrec_S 1 (λ_ r ⇒ x * r ) n Hn .
∎
Proof: Let x be given.
Assume Hx .
We will
prove x * x ^ 0 = x .
rewrite the current goal using
exp_SNo_nat_0 x Hx (from left to right).
An
exact proof term for the current goal is
mul_SNo_oneR x Hx .
∎
Proof: Let x be given.
Assume Hx .
We will
prove x * x ^ 1 = x * x .
Use f_equal.
∎
Proof: Let x be given.
Assume Hx .
rewrite the current goal using
exp_SNo_nat_2 x Hx (from left to right).
∎
Proof: Let x be given.
Assume Hx .
rewrite the current goal using
exp_SNo_nat_2 x Hx (from left to right).
∎
Proof: Let x be given.
Assume Hx .
We will
prove SNo (x ^ 0 ) .
rewrite the current goal using
exp_SNo_nat_0 x Hx (from left to right).
An
exact proof term for the current goal is
SNo_1 .
Let n be given.
Assume Hn .
rewrite the current goal using
exp_SNo_nat_S x Hx n Hn (from left to right).
We will
prove SNo (x * x ^ n ) .
An
exact proof term for the current goal is
SNo_mul_SNo x (x ^ n ) Hx IHn .
∎
Proof: Let x be given.
Assume Hx .
We prove the intermediate
claim Lx :
SNo x .
An
exact proof term for the current goal is
nat_p_SNo x Hx .
rewrite the current goal using
exp_SNo_nat_0 x Lx (from left to right).
An
exact proof term for the current goal is
nat_1 .
Let n be given.
Assume Hn .
rewrite the current goal using
exp_SNo_nat_S x Lx n Hn (from left to right).
An
exact proof term for the current goal is
mul_nat_p x Hx (x ^ n ) IHn .
∎
Proof:
Let n be given.
Assume Hn .
We prove the intermediate
claim Ln :
n ∈ ω .
An
exact proof term for the current goal is
nat_p_omega n Hn .
rewrite the current goal using
eps_SNoCut n Ln (from left to right).
We prove the intermediate
claim Lx :
SNo x .
We prove the intermediate
claim Lx2 :
0 < x .
rewrite the current goal using
add_SNo_eq x Lx x Lx (from left to right).
Let w be given.
Apply SNoL_E x Lx w Hw to the current goal.
We prove the intermediate
claim Lw0 :
w ≤ 0 .
Apply SNoLt_tra w x w Hw1 Lx Hw1 Hw3 to the current goal.
An
exact proof term for the current goal is
SNoLev_ w Hw1 .
An exact proof term for the current goal is H1 .
An exact proof term for the current goal is H1 .
We prove the intermediate
claim L1a :
w + x < eps_ n .
An
exact proof term for the current goal is
SNo_add_SNo w x Hw1 Lx .
An
exact proof term for the current goal is
SNo_eps_ n Ln .
We will
prove w + x ≤ 0 + x .
rewrite the current goal using
add_SNo_0L x Lx (from left to right).
Apply andI to the current goal.
An exact proof term for the current goal is L1a .
rewrite the current goal using
add_SNo_com x w Lx Hw1 (from left to right).
An exact proof term for the current goal is L1a .
Let w' be given.
rewrite the current goal using
eps_SNoCut n Ln (from right to left).
We will
prove w' < eps_ n .
Let w be given.
rewrite the current goal using Hw2 (from left to right).
An
exact proof term for the current goal is
andEL (w + x < eps_ n ) (x + w < eps_ n ) (L1 w Hw1 ) .
Let w be given.
rewrite the current goal using Hw2 (from left to right).
An
exact proof term for the current goal is
andER (w + x < eps_ n ) (x + w < eps_ n ) (L1 w Hw1 ) .
Let z be given.
Apply SNoR_E x Lx z Hz to the current goal.
We prove the intermediate
claim L2a :
eps_ n < z + x .
An
exact proof term for the current goal is
SNo_add_SNo z x Hz1 Lx .
We will
prove eps_ n ≤ z .
An
exact proof term for the current goal is
SNoLev_ z Hz1 .
An exact proof term for the current goal is Hz3 .
rewrite the current goal using
add_SNo_0R z Hz1 (from right to left) at position 1.
We will
prove z + 0 < z + x .
Apply andI to the current goal.
An exact proof term for the current goal is L2a .
rewrite the current goal using
add_SNo_com x z Lx Hz1 (from left to right).
An exact proof term for the current goal is L2a .
Let z' be given.
rewrite the current goal using
eps_SNoCut n Ln (from right to left).
We will
prove eps_ n < z' .
Let z be given.
rewrite the current goal using Hz2 (from left to right).
An
exact proof term for the current goal is
andEL (eps_ n < z + x ) (eps_ n < x + z ) (L2 z Hz1 ) .
Let z be given.
rewrite the current goal using Hz2 (from left to right).
An
exact proof term for the current goal is
andER (eps_ n < z + x ) (eps_ n < x + z ) (L2 z Hz1 ) .
Let w be given.
rewrite the current goal using
SingE 0 w Hw (from left to right).
rewrite the current goal using
add_SNo_eq x Lx x Lx (from right to left).
We will
prove 0 + 0 < x + x .
Let z be given.
rewrite the current goal using
add_SNo_eq x Lx x Lx (from right to left).
Let m be given.
rewrite the current goal using Hm2 (from left to right).
rewrite the current goal using IH m Hm1 (from right to left).
We will
prove x + x < y + y .
We prove the intermediate
claim Lm :
m ∈ ω .
An
exact proof term for the current goal is
nat_p_trans n Hn m Hm1 .
We prove the intermediate
claim Ly :
SNo y .
We prove the intermediate
claim Lxy :
x < y .
An exact proof term for the current goal is Hm1 .
An
exact proof term for the current goal is
add_SNo_Lt3 x x y y Lx Lx Ly Ly Lxy Lxy .
∎
Proof: Use transitivity with and
eps_ 0 .
An
exact proof term for the current goal is
eps_0_1 .
∎
Proof: rewrite the current goal using
add_SNo_1_1_2 (from right to left).
∎
Proof: Let x, y and z be given.
Assume Hx Hy Hz H1 .
We prove the intermediate
claim Lyz :
SNo (y + z ) .
An
exact proof term for the current goal is
SNo_add_SNo y z Hy Hz .
We will
prove 2 * x = 1 * (y + z ) .
rewrite the current goal using
mul_SNo_oneL (y + z ) Lyz (from left to right).
We will
prove 2 * x = y + z .
rewrite the current goal using
add_SNo_1_1_2 (from right to left).
We will
prove (1 + 1 ) * x = y + z .
We will
prove 1 * x + 1 * x = y + z .
rewrite the current goal using
mul_SNo_oneL x Hx (from left to right).
We will
prove x + x = y + z .
An exact proof term for the current goal is H1 .
∎
Proof: Let x be given.
Assume Hx Hx1 .
rewrite the current goal using
exp_SNo_nat_0 x Hx (from left to right).
Let n be given.
Assume Hn .
We prove the intermediate
claim Lxn :
SNo (x ^ n ) .
rewrite the current goal using
exp_SNo_nat_S x Hx n Hn (from left to right).
We will
prove 1 ≤ x * x ^ n .
We will
prove 1 * 1 ≤ x * x ^ n .
An
exact proof term for the current goal is
SNoLt_0_1 .
An
exact proof term for the current goal is
SNoLt_0_1 .
An exact proof term for the current goal is Hx1 .
An exact proof term for the current goal is IHn .
∎
Proof:
An
exact proof term for the current goal is
SNoLt_0_1 .
Let n be given.
Assume Hn .
We prove the intermediate
claim Ln :
SNo n .
An
exact proof term for the current goal is
nat_p_SNo n Hn .
We prove the intermediate
claim L2n :
SNo (2 ^ n ) .
rewrite the current goal using
add_SNo_1_1_2 (from right to left) at position 1.
rewrite the current goal using
mul_SNo_oneL (2 ^ n ) L2n (from left to right).
An
exact proof term for the current goal is
SNoLt_1_2 .
An exact proof term for the current goal is Hn .
∎
Proof:
rewrite the current goal using
eps_0_1 (from left to right).
Let n be given.
Assume Hn .
rewrite the current goal using
add_SNo_1_1_2 (from right to left) at position 1.
An exact proof term for the current goal is IHn .
∎
Theorem. (
eps_bd_1 ) The following is provable:
Proof: Let n be given.
Assume Hn .
An
exact proof term for the current goal is
SNo_eps_pos n Hn .
An
exact proof term for the current goal is
SNo_1 .
We will
prove SNo (2 ^ n ) .
An
exact proof term for the current goal is
SNo_2 .
An
exact proof term for the current goal is
omega_nat_p n Hn .
An
exact proof term for the current goal is
SNoLt_1_2 .
An
exact proof term for the current goal is
omega_nat_p n Hn .
∎
Proof: Let n be given.
Assume Hn .
Use transitivity with and
eps_ n * 2 ^ n .
∎
Proof: Let x be given.
Assume Hx .
Let m be given.
Assume Hm .
We prove the intermediate
claim Lm :
SNo m .
An
exact proof term for the current goal is
nat_p_SNo m Hm .
We will
prove x ^ m * x ^ 0 = x ^ (m + 0 ) .
rewrite the current goal using
exp_SNo_nat_0 x Hx (from left to right).
rewrite the current goal using
add_SNo_0R m Lm (from left to right).
Let n be given.
rewrite the current goal using
exp_SNo_nat_S x Hx n Hn (from left to right).
rewrite the current goal using
add_nat_SR m n Hn (from left to right).
We will
prove x ^ m * (x * x ^ n ) = x * x ^ (m + n ) .
rewrite the current goal using IHn (from right to left).
We will
prove x ^ m * (x * x ^ n ) = x * (x ^ m * x ^ n ) .
We will
prove (x ^ m * x ) * x ^ n = x * (x ^ m * x ^ n ) .
We will
prove (x * x ^ m ) * x ^ n = x * (x ^ m * x ^ n ) .
Use symmetry.
∎
Proof: Let x be given.
Assume Hx .
Let m be given.
Assume Hm .
Let n be given.
Assume Hn .
∎
Proof: Let x be given.
Assume Hx Hxpos .
rewrite the current goal using
exp_SNo_nat_0 x Hx (from left to right).
An
exact proof term for the current goal is
SNoLt_0_1 .
Let n be given.
rewrite the current goal using
exp_SNo_nat_S x Hx n Hn (from left to right).
We will
prove 0 < x * x ^ n .
rewrite the current goal using
mul_SNo_zeroR x Hx (from right to left).
We will
prove x * 0 < x * x ^ n .
∎
Proof: Let m be given.
Assume Hm .
Let n be given.
Assume Hn .
We prove the intermediate
claim Lmn1 :
m + n ∈ ω .
We prove the intermediate
claim Lmn2 :
nat_p (m + n ) .
An
exact proof term for the current goal is
omega_nat_p (m + n ) Lmn1 .
We prove the intermediate
claim Lem :
SNo (eps_ m ) .
An
exact proof term for the current goal is
SNo_eps_ m Hm .
We prove the intermediate
claim Len :
SNo (eps_ n ) .
An
exact proof term for the current goal is
SNo_eps_ n Hn .
We prove the intermediate
claim Lemen :
SNo (eps_ m * eps_ n ) .
We prove the intermediate
claim Lemn :
SNo (eps_ (m + n ) ) .
An
exact proof term for the current goal is
SNo_eps_ (m + n ) Lmn1 .
We prove the intermediate
claim L2m :
SNo (2 ^ m ) .
We prove the intermediate
claim L2n :
SNo (2 ^ n ) .
We will
prove SNo (2 ^ (m + n ) ) .
We will
prove 2 ^ (m + n ) ≠ 0 .
rewrite the current goal using H1 (from right to left) at position 2.
We will
prove 0 < 2 ^ (m + n ) .
An exact proof term for the current goal is Lemen .
An exact proof term for the current goal is Lemn .
Use f_equal.
Use symmetry.
∎
Proof:
Set f to be the term
λi : set ⇒ 0 .
We use f to witness the existential quantifier.
Apply bijI to the current goal.
Let i be given.
Apply SepI to the current goal.
An
exact proof term for the current goal is
SNoLev_0 .
Let i be given.
Let j be given.
Apply cases_1 i Hi to the current goal.
Apply cases_1 j Hj to the current goal.
An exact proof term for the current goal is Hij .
Let x be given.
We prove the intermediate
claim L1 :
x = 0 .
Use symmetry.
rewrite the current goal using
SNoLev_0 (from left to right).
Use symmetry.
An exact proof term for the current goal is Hx2 .
rewrite the current goal using
SNoLev_0 (from left to right).
Let alpha be given.
An
exact proof term for the current goal is
EmptyE alpha Ha .
We use
0 to
witness the existential quantifier.
Apply andI to the current goal.
An
exact proof term for the current goal is
In_0_1 .
Use symmetry.
An exact proof term for the current goal is L1 .
Let n be given.
Assume Hn .
rewrite the current goal using
add_SNo_1_1_2 (from right to left) at position 1.
We prove the intermediate
claim L2n1 :
nat_p (2 ^ n ) .
We prove the intermediate
claim L2n2 :
ordinal (2 ^ n ) .
An exact proof term for the current goal is L2n1 .
We prove the intermediate
claim L2n3 :
SNo (2 ^ n ) .
An exact proof term for the current goal is L2n2 .
We prove the intermediate
claim L2n2n1 :
nat_p (2 ^ n + 2 ^ n ) .
An exact proof term for the current goal is L2n1 .
An exact proof term for the current goal is L2n1 .
We prove the intermediate
claim L2n2n2 :
ordinal (2 ^ n + 2 ^ n ) .
An exact proof term for the current goal is L2n2n1 .
We prove the intermediate
claim L2n2n3 :
SNo (2 ^ n + 2 ^ n ) .
An exact proof term for the current goal is L2n2n2 .
We prove the intermediate
claim L2npLt2n2n :
∀m, SNo m → m < 2 ^ n → 2 ^ n + m < 2 ^ n + 2 ^ n .
Let m be given.
Assume Hm H1 .
An
exact proof term for the current goal is
add_SNo_Lt2 (2 ^ n ) m (2 ^ n ) L2n3 Hm L2n3 H1 .
We prove the intermediate
claim L2nLt2n2n :
2 ^ n < 2 ^ n + 2 ^ n .
rewrite the current goal using
add_SNo_0R (2 ^ n ) L2n3 (from right to left) at position 1.
Apply L2npLt2n2n 0 SNo_0 to the current goal.
Apply IHn to the current goal.
Let f be given.
Let x be given.
Assume Hx .
Assume Hx HxSn .
Let p be given.
Assume Hp .
We prove the intermediate
claim L2a :
n ∈ SNoLev x .
rewrite the current goal using HxSn (from left to right).
An
exact proof term for the current goal is
restr_SNoLev x Hx3 n L2a .
Apply Hp to the current goal.
An exact proof term for the current goal is Hx3 .
An exact proof term for the current goal is HxSn .
Apply SepI to the current goal.
An
exact proof term for the current goal is
restr_SNo_ x Hx3 n L2a .
An exact proof term for the current goal is L2b .
An
exact proof term for the current goal is
restr_SNo x Hx3 n L2a .
An exact proof term for the current goal is L2b .
Let u be given.
Assume Hu .
Let p be given.
Assume Hp .
Apply L2 u Hu to the current goal.
We prove the intermediate
claim Lf1a :
f (u ∩ SNoElts_ n ) ∈ 2 ^ n .
An
exact proof term for the current goal is
Hf1 (u ∩ SNoElts_ n ) Hu1 .
Apply Hp to the current goal.
An exact proof term for the current goal is Lf1b .
An exact proof term for the current goal is Lf1c .
An exact proof term for the current goal is Lf1c .
An exact proof term for the current goal is L2n2 .
An exact proof term for the current goal is Lf1a .
Apply andI to the current goal.
Apply Lg to the current goal.
Let g be given.
Assume H .
Apply H to the current goal.
We use g to witness the existential quantifier.
Apply bijI to the current goal.
Let u be given.
We will
prove g u ∈ 2 ^ n + 2 ^ n .
Apply L2 u Hu to the current goal.
Apply Lf u Hu to the current goal.
Apply xm (n ∈ u ) to the current goal.
rewrite the current goal using Hg1 u H1 (from left to right).
An
exact proof term for the current goal is
SNoLt_tra (f (u ∩ SNoElts_ n ) ) (2 ^ n ) (2 ^ n + 2 ^ n ) Lfu3 L2n3 L2n2n3 Lfu4 L2nLt2n2n .
rewrite the current goal using Hg2 u H1 (from left to right).
An
exact proof term for the current goal is
L2npLt2n2n (f (u ∩ SNoElts_ n ) ) Lfu3 Lfu4 .
Let u be given.
Let v be given.
Apply L2 u Hu to the current goal.
Apply Lf u Hu to the current goal.
Apply L2 v Hv to the current goal.
Apply Lf v Hv to the current goal.
We prove the intermediate
claim LnLu :
n ∈ SNoLev u .
rewrite the current goal using Hu0b (from left to right).
We prove the intermediate
claim LnLv :
n ∈ SNoLev v .
rewrite the current goal using Hv0b (from left to right).
Apply xm (n ∈ u ) to the current goal.
rewrite the current goal using Hg1 u H1 (from left to right).
Apply xm (n ∈ v ) to the current goal.
rewrite the current goal using Hg1 v H2 (from left to right).
An
exact proof term for the current goal is
Hf2 (u ∩ SNoElts_ n ) Hu1 (v ∩ SNoElts_ n ) Hv1 Hguv .
Apply SNo_eq u v Hu0a Hv0a to the current goal.
rewrite the current goal using Hv0b (from left to right).
An exact proof term for the current goal is Hu0b .
rewrite the current goal using Hu0b (from left to right).
Let i be given.
Apply ordsuccE n i Hi to the current goal.
We will
prove i ∈ u ↔ i ∈ v .
An
exact proof term for the current goal is
restr_SNoEq u ?? n LnLu i H3 .
rewrite the current goal using Luv (from left to right).
An
exact proof term for the current goal is
restr_SNoEq v ?? n LnLv i H3 .
rewrite the current goal using H3 (from left to right).
Apply iffI to the current goal.
Assume _ .
An exact proof term for the current goal is H2 .
Assume _ .
An exact proof term for the current goal is H1 .
rewrite the current goal using Hg2 v H2 (from left to right).
We will
prove 2 ^ n < 2 ^ n .
rewrite the current goal using
add_SNo_0R (2 ^ n ) L2n3 (from right to left) at position 1.
rewrite the current goal using Hguv (from right to left).
An exact proof term for the current goal is Lfu4 .
rewrite the current goal using Hg2 u H1 (from left to right).
Apply xm (n ∈ v ) to the current goal.
rewrite the current goal using Hg1 v H2 (from left to right).
We will
prove 2 ^ n < 2 ^ n .
rewrite the current goal using
add_SNo_0R (2 ^ n ) L2n3 (from right to left) at position 1.
rewrite the current goal using Hguv (from left to right).
An exact proof term for the current goal is Lfv4 .
rewrite the current goal using Hg2 v H2 (from left to right).
Apply SNo_eq u v Hu0a Hv0a to the current goal.
rewrite the current goal using Hv0b (from left to right).
An exact proof term for the current goal is Hu0b .
rewrite the current goal using Hu0b (from left to right).
Let i be given.
Apply ordsuccE n i Hi to the current goal.
We will
prove i ∈ u ↔ i ∈ v .
An
exact proof term for the current goal is
restr_SNoEq u ?? n LnLu i H3 .
rewrite the current goal using Luv (from left to right).
An
exact proof term for the current goal is
restr_SNoEq v ?? n LnLv i H3 .
rewrite the current goal using H3 (from left to right).
Apply iffI to the current goal.
An exact proof term for the current goal is H1 H4 .
An exact proof term for the current goal is H2 H4 .
Let m be given.
We prove the intermediate
claim Lm1 :
nat_p m .
An
exact proof term for the current goal is
nat_p_trans (2 ^ n + 2 ^ n ) L2n2n1 m Hm .
We prove the intermediate
claim Lm2 :
SNo m .
An
exact proof term for the current goal is
nat_p_SNo m Lm1 .
Apply Hf3 m H1 to the current goal.
Let u be given.
Assume H .
Apply H to the current goal.
Assume Hu3a Hu3b Hu3c Hu3d .
rewrite the current goal using Hu4 (from right to left).
rewrite the current goal using Hu4 (from right to left).
We use
(SNo_extend1 u ) to
witness the existential quantifier.
Apply andI to the current goal.
Apply SepI to the current goal.
An
exact proof term for the current goal is
nat_p_omega n Hn .
rewrite the current goal using Hu4 (from right to left).
An exact proof term for the current goal is Lu2 .
rewrite the current goal using
Hg1 (SNo_extend1 u ) Lu3 (from left to right).
rewrite the current goal using Hu4 (from right to left).
An exact proof term for the current goal is Hu2 .
Apply Hf3 (m + - 2 ^ n ) H1 to the current goal.
Let u be given.
Assume H .
Apply H to the current goal.
Assume Hu3a Hu3b Hu3c Hu3d .
rewrite the current goal using Hu4 (from right to left).
rewrite the current goal using Hu4 (from right to left).
We use
(SNo_extend0 u ) to
witness the existential quantifier.
Apply andI to the current goal.
Apply SepI to the current goal.
An
exact proof term for the current goal is
nat_p_omega n Hn .
rewrite the current goal using Hu4 (from right to left).
An exact proof term for the current goal is Lu2 .
rewrite the current goal using
Hg2 (SNo_extend0 u ) Lu3 (from left to right).
rewrite the current goal using Hu4 (from right to left) at position 2.
We will
prove 2 ^ n + f u = m .
rewrite the current goal using Hu2 (from left to right).
We will
prove 2 ^ n + (m + - 2 ^ n ) = m .
We will
prove (m + - 2 ^ n ) + 2 ^ n = m .
We will
prove m + (- 2 ^ n + 2 ^ n ) = m .
An
exact proof term for the current goal is
add_SNo_0R m Lm2 .
∎
Proof: Let n be given.
Assume Hn .
We prove the intermediate
claim Ln :
nat_p n .
An
exact proof term for the current goal is
omega_nat_p n Hn .
We prove the intermediate
claim Ln2 :
ordinal n .
Let x be given.
Assume Hx .
Apply SNoS_E2 n Ln2 x Hx to the current goal.
Apply SepI to the current goal.
An exact proof term for the current goal is Hx .
Let x be given.
Assume Hx .
Let i be given.
Assume H1 .
rewrite the current goal using Hxb (from left to right).
An exact proof term for the current goal is Hi .
An exact proof term for the current goal is Hx4 .
rewrite the current goal using L1 (from left to right).
Let i be given.
We use
2 ^ i to
witness the existential quantifier.
Apply andI to the current goal.
∎
Proof: Let x be given.
Assume Hx .
Let y be given.
Assume Hy .
Apply SNoL_E x Hx3 y Hy to the current goal.
Assume Hy1 Hy2 Hy3 .
Apply SNoS_I2 y x ?? ?? to the current goal.
An exact proof term for the current goal is Hy2 .
∎
Proof: Let x be given.
Assume Hx .
Let y be given.
Assume Hy .
Apply SNoR_E x Hx3 y Hy to the current goal.
Assume Hy1 Hy2 Hy3 .
Apply SNoS_I2 y x ?? ?? to the current goal.
An exact proof term for the current goal is Hy2 .
∎
End of Section SurrealExp