Built with Alectryon, running Coq+SerAPI v8.19.0+0.19.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
[Loading ML file number_string_notation_plugin.cmxs (using legacy method) ... done]
Require Import TruncType HSet.Require Import HoTT.Truncations.Core.LocalOpen Scope nat_scope.LocalOpen Scope path_scope.(** * The surreal numbers *)(** Based on section 11.6 of the HoTT Book. *)Declare Scope surreal_scope.Delimit Scope surreal_scope with No.LocalOpen Scope surreal_scope.(** ** Option sorts *)(** We refine the surreal numbers by parametrizing them by "option sorts", which are predicates on the types that index the options. A surreal number with given option sorts "hereditarily" has all options parametrized by types belonging to that sort. *)DefinitionOptionSort@{i} := Type@{i} -> Type@{i} -> Type@{i}.ClassInSort (S : OptionSort@{i}) (IJ : Type@{i})
:= insort : S I J.(** The surreal numbers use a lot of universes. We include some universe annotations here and there to reduce the number of overall universe parameters from an unmanageable number to a slightly less unmanageable number. This improves performance significantly. We also use [abstract] and [Qed] whenever possible, for the same reason. *)(** ** Definition *)ModuleExport Surreals.SectionOptionSort.(** We will use this to assert that certain inequalities below hold. We locate it here, so that it depends on no universe variables. See the longer explanation below. *)InductiveNo_Empty_for_admitted : Type0 := .AxiomNo_Empty_admitted : No_Empty_for_admitted.Universei.Context {S : OptionSort@{i}}.(** *** Games first *)(** Since Coq doesn't support inductive-inductive types natively, we have to hack a bit. Inspired by Conway, we define [Game]s to be constructed by the point-constructor of [No] but without the hypothesis on inequality of options. Then we define the inequalities as a mutual inductive family over [Game], and put an inductive predicate on [Game] characterizing those that are Numbers. (This is roughly a standard technique described by Fredrik Nordvall Forsberg for reducing induction-induction to parametrized induction.) We then proceed to add axioms for the path-constructors of [No]. It should be emphasized that this is *not* currently a correct higher inductive-inductive definition of games; these "games" are only being used inside this module as a trick to produce [No] in a way that computes on the point-constructor. It should be possible to make a higher inductive-inductive definition of games, but this is not it. *)Private InductiveGame : Type :=
| opt : forall (LR : Type@{i})
(s : InSort S L R)
(xL : L -> Game) (xR : R -> Game), Game.Arguments opt {L R s} xL xR.Private Inductivegame_le : Game -> Game -> Type :=
| game_le_lr
: forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> Game) (xR : R -> Game)
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> Game) (yR : R' -> Game),
(forall (l:L), game_lt (xL l) (opt yL yR)) ->
(forall (r:R'), game_lt (opt xL xR) (yR r)) ->
game_le (opt xL xR) (opt yL yR)
with game_lt : Game -> Game -> Type :=
| game_lt_l
: forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> Game) (xR : R -> Game)
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> Game) (yR : R' -> Game)
(l : L'),
(game_le (opt xL xR) (yL l)) ->
game_lt (opt xL xR) (opt yL yR)
| game_lt_r
: forall (LR : Type@{i}) (s : InSort@{i} S L R)
(xL : L -> Game) (xR : R -> Game)
(L'R' : Type@{i}) (s' : InSort@{i} S L' R')
(yL : L' -> Game) (yR : R' -> Game)
(r : R),
(game_le (xR r) (opt yL yR)) ->
game_lt (opt xL xR) (opt yL yR).Arguments game_le_lr {L R s} xL xR {L' R' s'} yL yR _ _.Arguments game_lt_l {L R s} xL xR {L' R' s'} yL yR l _.Arguments game_lt_r {L R s} xL xR {L' R' s'} yL yR r _.(** *** Now the surreals *)Private Inductiveis_surreal : Game -> Type :=
| isno : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> Game) (xR : R -> Game),
(foralll, is_surreal (xL l))
-> (forallr, is_surreal (xR r))
-> (forall (l:L) (r:R), game_lt (xL l) (xR r))
-> is_surreal (opt xL xR).Unset Nonrecursive Elimination Schemes.(** We call these "general surreal numbers" since they are parametrized by an option sort. *)RecordGenNo : Type := Build_No
{ game_of : Game
; isno_game_of : is_surreal (game_of) }.Bind Scope surreal_scope with GenNo.Definitionlt (xy : GenNo) := game_lt (game_of x) (game_of y).Definitionle (xy : GenNo) := game_le (game_of x) (game_of y).LocalInfix"<" := lt : surreal_scope.LocalInfix"<=" := le : surreal_scope.DefinitionNo_cut {LR : Type@{i}} {s : InSort S L R}
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
: GenNo
:= Build_No (opt (game_of o xL) (game_of o xR))
(isno _ _ _ _ _ (isno_game_of o xL)
(isno_game_of o xR) xcut).Notation"{ { xL | xR // xcut } }" := (No_cut xL xR xcut) : surreal_scope.Axiompath_No : forall (xy : GenNo), (x <= y) -> (y <= x) -> (x = y).Arguments path_No {x y} _ _.Definitionle_lr
{LR : Type@{i} } {s : InSort S L R}
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
{L'R' : Type@{i} } {s' : InSort S L' R'}
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
: (forall (l:L), xL l < {{ yL | yR // ycut }}) ->
(forall (r:R'), {{ xL | xR // xcut }} < yR r) ->
{{ xL | xR // xcut }} <= {{ yL | yR // ycut }}
:= game_le_lr (game_of o xL) (game_of o xR)
(game_of o yL) (game_of o yR).Definitionlt_l
{LR : Type@{i} } {s : InSort S L R}
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
{L'R' : Type@{i} } {s' : InSort S L' R'}
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(l : L')
: ({{ xL | xR // xcut }} <= yL l) ->
{{ xL | xR // xcut }} < {{ yL | yR // ycut }}
:= game_lt_l (game_of o xL) (game_of o xR)
(game_of o yL) (game_of o yR) l.Definitionlt_r
{LR : Type@{i} } {s : InSort S L R}
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
{L'R' : Type@{i} } {s' : InSort S L' R'}
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(r : R)
: (xR r <= {{ yL | yR // ycut }}) ->
{{ xL | xR // xcut }} < {{ yL | yR // ycut }}
:= game_lt_r (game_of o xL) (game_of o xR)
(game_of o yL) (game_of o yR) r.
S: OptionSort x, y: GenNo
IsHProp (x <= y)
Admitted.
S: OptionSort x, y: GenNo
IsHProp (x < y)
Admitted.(** *** Now the induction principle. *)SectionNoInd.Context
(A : GenNo -> Type)
(dle : forall (xy : GenNo), (x <= y) -> A x -> A y -> Type)
(dlt : forall (xy : GenNo), (x < y) -> A x -> A y -> Type)
{ishprop_le : forallxyabp, IsHProp (dle x y p a b)}
{ishprop_lt : forallxyabp, IsHProp (dlt x y p a b)}
(dcut : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : foralll, A (xL l)) (fxR : forallr, A (xR r))
(fxcut : foralllr, dlt _ _ (xcut l r) (fxL l) (fxR r)),
A {{ xL | xR // xcut }})
(dpath : forall (xy : GenNo) (a:A x) (b:A y) (p : x <= y) (q : y <= x)
(dp : dle x y p a b) (dq : dle y x q b a),
path_No p q # a = b)
(dle_lr : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : foralll, A (xL l)) (fxR : forallr, A (xR r))
(fxcut : foralllr, dlt _ _ (xcut l r) (fxL l) (fxR r))
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(fyL : foralll, A (yL l)) (fyR : forallr, A (yR r))
(fycut : foralllr, dlt _ _ (ycut l r) (fyL l) (fyR r))
(p : forall (l:L), xL l < {{ yL | yR // ycut }})
(dp : forall (l:L),
dlt _ _ (p l)
(fxL l)
(dcut _ _ _ yL yR ycut fyL fyR fycut))
(q : forall (r:R'), {{ xL | xR // xcut }} < yR r)
(dq : forall (r:R'),
dlt _ _ (q r)
(dcut _ _ _ xL xR xcut fxL fxR fxcut)
(fyR r)),
dle {{ xL | xR // xcut }} {{ yL | yR // ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut _ _ _ xL xR xcut fxL fxR fxcut)
(dcut _ _ _ yL yR ycut fyL fyR fycut))
(dlt_l : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : foralll, A (xL l)) (fxR : forallr, A (xR r))
(fxcut : foralllr, dlt _ _ (xcut l r) (fxL l) (fxR r))
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(fyL : foralll, A (yL l)) (fyR : forallr, A (yR r))
(fycut : foralllr, dlt _ _ (ycut l r) (fyL l) (fyR r))
(l : L')
(p : {{ xL | xR // xcut }} <= yL l)
(dp : dle _ _ p
(dcut _ _ _ xL xR xcut fxL fxR fxcut)
(fyL l)),
dlt {{ xL | xR // xcut }} {{ yL | yR // ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut _ _ _ xL xR xcut fxL fxR fxcut)
(dcut _ _ _ yL yR ycut fyL fyR fycut))
(dlt_r : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : foralll, A (xL l)) (fxR : forallr, A (xR r))
(fxcut : foralllr, dlt _ _ (xcut l r) (fxL l) (fxR r))
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(fyL : foralll, A (yL l)) (fyR : forallr, A (yR r))
(fycut : foralllr, dlt _ _ (ycut l r) (fyL l) (fyR r))
(r : R)
(p : (xR r) <= {{ yL | yR // ycut }})
(dp : dle _ _ p
(fxR r)
(dcut _ _ _ yL yR ycut fyL fyR fycut)),
dlt {{ xL | xR // xcut }} {{ yL | yR // ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut _ _ _ xL xR xcut fxL fxR fxcut)
(dcut _ _ _ yL yR ycut fyL fyR fycut)).(** As usual for HITs implemented with [Private Inductive], we will define [No_ind] inside this module using a fixpoint over [No], thereby obtaining a judgmental computation rule for the point-constructor [No_cut], and then assert the other computation rules as axioms. In this case, the relevant other rules are the preservation of inequalities.However, it turns out that in defining [No_cut] we already need to know that it preserves inequalities. Since this is eventually an axiom anyway, we could just assert it with [admit] in the proof. However, if we did this then the [admit] would not be *judgmentally* equal to the axiom [No_ind_lt] that we assert afterwards. Instead, we make use of the fact that [admit] is essentially by definition [match proof_admitted with end] for a global axiom [proof_admitted : Empty], so that if we use the same [admit] both inside the definition of [No_ind] and in asserting [No_ind_lt] as an axiom, they will be the same term judgmentally.Finally, for conceptual isolation, and so as not to depend on the particular implementation of [admit], we use local copies of [Empty] and [proof_admitted]. These were defined at the start of the Section, because otherwise they depend on six universe variables. *)(** Technically, we induct over the inductive predicate witnessing Numberhood of games. We prove the "induction step" separately to improve performance, possibly by preventing bare [fix]s from appearing upon simplification. *)
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }}
(yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} x: Game xno: is_surreal x
A {| game_of := x; isno_game_of := xno |}
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }}
(yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} x: Game xno: is_surreal x
A {| game_of := x; isno_game_of := xno |}
(* We use [revert] and [intros] as a way to ensure that the definition depends on all of the variables in the context. *)
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} x: Game xno: is_surreal x
(forall (xy : GenNo) (a : A x) (b : A y)
(p : x <= y), IsHProp (dle x y p a b)) ->
(forall (xy : GenNo) (a : A x) (b : A y) (p : x < y),
IsHProp (dlt x y p a b)) ->
(forall (xy : GenNo) (a : A x) (b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a -> transport A (path_No p q) a = b) ->
(forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L') (r : R'),
dlt (yL l) (yR r) (ycut l r) (fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }} (p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }} (yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut) (fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L') (r : R'),
dlt (yL l) (yR r) (ycut l r) (fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }} (yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut) (fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L') (r : R'),
dlt (yL l) (yR r) (ycut l r) (fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
A {| game_of := x; isno_game_of := xno |}
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} L, R: Type s: InSort S L R xL: L -> Game xR: R -> Game Lno: foralll : L, is_surreal (xL l) Rno: forallr : R, is_surreal (xR r) xcut: forall (l : L) (r : R), game_lt (xL l) (xR r)
(forall (xy : GenNo) (a : A x) (b : A y)
(p : x <= y), IsHProp (dle x y p a b)) ->
(forall (xy : GenNo) (a : A x) (b : A y) (p : x < y),
IsHProp (dlt x y p a b)) ->
(forall (xy : GenNo) (a : A x) (b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a -> transport A (path_No p q) a = b) ->
(forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L') (r : R'),
dlt (yL l) (yR r) (ycut l r) (fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }} (p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }} (yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut) (fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L') (r : R'),
dlt (yL l) (yR r) (ycut l r) (fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }} (yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut) (fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L') (r : R'),
dlt (yL l) (yR r) (ycut l r) (fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
A
{|
game_of := opt xL xR;
isno_game_of := isno L R s xL xR Lno Rno xcut
|}
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} L, R: Type s: InSort S L R xL: L -> Game xR: R -> Game Lno: foralll : L, is_surreal (xL l) Rno: forallr : R, is_surreal (xR r) xcut: forall (l : L) (r : R), game_lt (xL l) (xR r) ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y) (q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (l : L')
(p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }} (yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (r : R)
(p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)
A
{|
game_of := opt xL xR;
isno_game_of := isno L R s xL xR Lno Rno xcut
|}
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} L, R: Type s: InSort S L R xL: L -> Game xR: R -> Game Lno: foralll : L, is_surreal (xL l) Rno: forallr : R, is_surreal (xR r) xcut: forall (l : L) (r : R), game_lt (xL l) (xR r) ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y) (q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (l : L')
(p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }} (yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (r : R)
(p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)
foralll : L,
A
((funl0 : L =>
{| game_of := xL l0; isno_game_of := Lno l0 |}) l)
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} L, R: Type s: InSort S L R xL: L -> Game xR: R -> Game Lno: foralll : L, is_surreal (xL l) Rno: forallr : R, is_surreal (xR r) xcut: forall (l : L) (r : R), game_lt (xL l) (xR r) ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y) (q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (l : L')
(p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }} (yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (r : R)
(p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)
forallr : R,
A
((funr0 : R =>
{| game_of := xR r0; isno_game_of := Rno r0 |}) r)
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} L, R: Type s: InSort S L R xL: L -> Game xR: R -> Game Lno: foralll : L, is_surreal (xL l) Rno: forallr : R, is_surreal (xR r) xcut: forall (l : L) (r : R), game_lt (xL l) (xR r) ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y) (q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (l : L')
(p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }} (yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L) (r : R),
dlt (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo)
(yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r)) (r : R)
(p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }}
(yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} x: Game xno: is_surreal x
A {| game_of := x; isno_game_of := xno |}
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }}
(yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) No_ind_internal: forall (x : Game)
(xno : is_surreal x),
A
{|
game_of := x;
isno_game_of := xno
|} x: Game xno: is_surreal x
A {| game_of := x; isno_game_of := xno |}
exact (No_ind_internal_step No_ind_internal x xno).Defined.
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }}
(yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x: GenNo
A x
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }}
(yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x: GenNo
A x
S: OptionSort A: GenNo -> Type dle: forallxy : GenNo, x <= y -> A x -> A y -> Type dlt: forallxy : GenNo, x < y -> A x -> A y -> Type ishprop_le: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y),
IsHProp (dle x y p a b) ishprop_lt: forall (xy : GenNo) (a : A x)
(b : A y) (p : x < y),
IsHProp (dlt x y p a b) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r)),
(forall (l : L) (r : R),
dlt (xL l) (xR r) (xcut l r) (fxL l) (fxR r)) ->
A {{ xL | xR//xcut }} dpath: forall (xy : GenNo) (a : A x)
(b : A y) (p : x <= y)
(q : y <= x),
dle x y p a b ->
dle y x q b a ->
transport A (path_No p q) a = b dle_lr: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(p : foralll : L, xL l < {{ yL | yR//ycut }}),
(foralll : L,
dlt (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
forallq : forallr : R', {{ xL | xR//xcut }} < yR r,
(forallr : R',
dlt {{ xL | xR//xcut }}
(yR r) (q r)
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(l : L') (p : {{ xL | xR//xcut }} <= yL l),
dle {{ xL | xR//xcut }}
(yL l) p
(dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L, A (xL l))
(fxR : forallr : R, A (xR r))
(fxcut : forall (l : L)
(r : R),
dlt (xL l) (xR r)
(xcut l r)
(fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo) (yR : R' -> GenNo)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : foralll : L', A (yL l))
(fyR : forallr : R', A (yR r))
(fycut : forall (l : L')
(r : R'),
dlt (yL l) (yR r)
(ycut l r)
(fyL l) (fyR r))
(r : R) (p : xR r <= {{ yL | yR//ycut }}),
dle (xR r) {{ yL | yR//ycut }} p
(fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
(dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x: Game xno: is_surreal x
A {| game_of := x; isno_game_of := xno |}
exact (No_ind_internal x xno).Defined.DefinitionNo_ind_le (xy : GenNo) (p : x <= y)
: dle x y p (No_ind x) (No_ind y)
:= match No_Empty_admitted withend.DefinitionNo_ind_lt (xy : GenNo) (p : x < y)
: dlt x y p (No_ind x) (No_ind y)
:= match No_Empty_admitted withend.(** Sometimes it's convenient to have all three parts of [No_ind] in one package, so that we only have to verify the hypotheses once. *)DefinitionNo_ind_package
: { f : forallx, A x &
(forall (xy : GenNo) (p : x <= y), dle x y p (f x) (f y)) *
(forall (xy : GenNo) (p : x < y), dlt x y p (f x) (f y)) }
:= ( No_ind ; (No_ind_le , No_ind_lt) ).(** It's also sometimes convenient to have just the inequality parts together. *)DefinitionNo_ind_le_lt (xy : GenNo)
: (forall (p : x <= y), dle x y p (No_ind x) (No_ind y)) *
(forall (p : x < y), dlt x y p (No_ind x) (No_ind y))
:= (No_ind_le x y , No_ind_lt x y).(** We verify that our definition computes judgmentally. *)DefinitionNo_ind_cut
(LR : Type@{i}) (s : InSort S L R)
(xL : L -> GenNo) (xR : R -> GenNo)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
: No_ind {{ xL | xR // xcut }}
= dcut L R _ xL xR xcut
(funl => No_ind (xL l)) (funr => No_ind (xR r))
(funlr => No_ind_lt (xL l) (xR r) (xcut l r))
:= 1.EndNoInd.EndOptionSort.Arguments GenNo S : clear implicits.Infix"<" := lt : surreal_scope.Infix"<=" := le : surreal_scope.Notation"{ { xL | xR // xcut } }" := (No_cut xL xR xcut) : surreal_scope.EndSurreals.SectionOptionSort.Universei.Context {S : OptionSort@{i}}.LetNo := GenNo S.(** ** A few surreal numbers *)Definitionrempty_cut {L : Type} {xL : L -> No}
: forall (l:L) (r:Empty), xL l < Empty_rec r
:= funl => Empty_ind _.Definitionlempty_cut {R : Type} {xR : R -> No}
: forall (l:Empty) (r:R), Empty_rec l < xR r
:= Empty_ind _.Definitionzero `{InSort S Empty Empty} : No
:= {{ Empty_rec | Empty_rec // lempty_cut }}.Definitionone `{InSort S Empty Empty} `{InSort S Unit Empty} : No
:= {{ unit_name zero | Empty_rec // rempty_cut }}.Definitionminusone `{InSort S Empty Empty} `{InSort S Empty Unit} : No
:= {{ Empty_rec | unit_name zero // lempty_cut }}.(** ** More induction principles *)(** *** The simplified induction principle for hprops *)
S: OptionSort No:= GenNo S: Type P: No -> Type H: forallx : No, IsHProp (P x) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r),
(foralll : L, P (xL l)) ->
(forallr : R, P (xR r)) ->
P {{ xL | xR//xcut }} x: No
P x
S: OptionSort No:= GenNo S: Type P: No -> Type H: forallx : No, IsHProp (P x) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r),
(foralll : L, P (xL l)) ->
(forallr : R, P (xR r)) ->
P {{ xL | xR//xcut }} x: No
P x
S: OptionSort No:= GenNo S: Type P: No -> Type H: forallx : No, IsHProp (P x) dcut: forall (LR : Type) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r),
(foralll : L, P (xL l)) ->
(forallr : R, P (xR r)) ->
P {{ xL | xR//xcut }} L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, P (xL l) fxR: forallr : R, P (xR r) fxcut: L -> R -> Unit
P {{ xL | xR//xcut }}
apply dcut; assumption.Defined.(** See also [repeat_No_ind_hprop], below *)(** *** The non-dependent recursion principle *)SectionNoRec.Context `{Funext}.Context (A : Type)
(dle : A -> A -> Type) `{is_mere_relation A dle}
(dlt : A -> A -> Type) `{is_mere_relation A dlt}
(dcut : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : L -> A) (fxR : R -> A)
(fxcut : foralllr, dlt (fxL l) (fxR r)),
A)
(dpath : forallab, dle a b -> dle b a -> a = b)
(dle_lr : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : L -> A) (fxR : R -> A)
(fxcut : foralllr, dlt (fxL l) (fxR r))
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(fyL : L' -> A) (fyR : R' -> A)
(fycut : foralllr, dlt (fyL l) (fyR r))
(p : forall (l:L), xL l < {{ yL | yR // ycut }})
(dp : forall (l:L),
dlt (fxL l) (dcut _ _ _ yL yR ycut fyL fyR fycut))
(q : forall (r:R'), {{ xL | xR // xcut }} < yR r)
(dq : forall (r:R'),
dlt (dcut _ _ _ xL xR xcut fxL fxR fxcut) (fyR r)),
dle (dcut _ _ _ xL xR xcut fxL fxR fxcut)
(dcut _ _ _ yL yR ycut fyL fyR fycut))
(dlt_l : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : L -> A) (fxR : R -> A)
(fxcut : foralllr, dlt (fxL l) (fxR r))
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(fyL : L' -> A) (fyR : R' -> A)
(fycut : foralllr, dlt (fyL l) (fyR r))
(l : L') (p : {{ xL | xR // xcut }} <= yL l)
(dp : dle (dcut _ _ _ xL xR xcut fxL fxR fxcut) (fyL l)),
dlt (dcut _ _ _ xL xR xcut fxL fxR fxcut)
(dcut _ _ _ yL yR ycut fyL fyR fycut))
(dlt_r : forall (LR : Type@{i}) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
(fxL : L -> A) (fxR : R -> A)
(fxcut : foralllr, dlt (fxL l) (fxR r))
(L'R' : Type@{i}) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
(fyL : L' -> A) (fyR : R' -> A)
(fycut : foralllr, dlt (fyL l) (fyR r))
(r : R) (p : (xR r) <= {{ yL | yR // ycut }})
(dp : dle (fxR r) (dcut _ _ _ yL yR ycut fyL fyR fycut)),
dlt (dcut _ _ _ xL xR xcut fxL fxR fxcut)
(dcut _ _ _ yL yR ycut fyL fyR fycut)).
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x: No
A
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x: No
A
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)
(fun_ : GenNo S => A) {{ xL | xR//xcut }}
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x, y: GenNo S a: (fun_ : GenNo S => A) x b: (fun_ : GenNo S => A) y p: x <= y q: y <= x dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) x y p a b dq: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) y x q b a
transport (fun_ : GenNo S => A) (path_No x y p q) a =
b
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) p: foralll : L, xL l < {{ yL | yR//ycut }} dp: foralll : L,
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall
(l0 : L)
(r : R),
xL l0 < xR r)
(fxL : foralll0 : L,
(fun_ : GenNo S => A) (xL l0))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l0 : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l0)
(xR r)
(xcut l0 r)
(fxL l0)
(fxR r)) =>
?Goal) L' R' s' yL yR ycut fyL fyR fycut) q: forallr : R', {{ xL | xR//xcut }} < yR r dq: forallr : R',
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) {{ xL | xR//xcut }}
(yR r) (q r)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall
(l : L)
(r0 : R),
xL l < xR r0)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr0 : R,
(fun_ : GenNo S => A) (xR r0))
(fxcut : forall
(l : L)
(r0 : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r0)
(xcut l r0)
(fxL l)
(fxR r0)) =>
?Goal) L R s xL xR xcut fxL fxR fxcut)
(fyR r)
(fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L R s xL xR xcut fxL fxR fxcut)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L' R' s' yL yR ycut fyL fyR fycut)
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) l: L' p: {{ xL | xR//xcut }} <= yL l dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) {{ xL | xR//xcut }}
(yL l) p
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L R s xL xR xcut fxL fxR fxcut)
(fyL l)
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L R s xL xR xcut fxL fxR fxcut)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L' R' s' yL yR ycut fyL fyR fycut)
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) r: R p: xR r <= {{ yL | yR//ycut }} dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) (xR r) {{ yL | yR//ycut }} p
(fxR r)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L' R' s' yL yR ycut fyL fyR fycut)
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) {{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L R s xL xR xcut fxL fxR fxcut)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
?Goal) L' R' s' yL yR ycut fyL fyR fycut)
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r)
(fun_ : GenNo S => A) {{ xL | xR//xcut }}
exact (dcut L R _ xL xR xcut fxL fxR fxcut).
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x, y: GenNo S a: (fun_ : GenNo S => A) x b: (fun_ : GenNo S => A) y p: x <= y q: y <= x dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) x y p a b dq: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) y x q b a
transport (fun_ : GenNo S => A) (path_No x y p q) a =
b
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) x, y: GenNo S a: (fun_ : GenNo S => A) x b: (fun_ : GenNo S => A) y p: x <= y q: y <= x dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) x y p a b dq: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) y x q b a
a = b
apply dpath; assumption.
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) p: foralll : L, xL l < {{ yL | yR//ycut }} dp: foralll : L,
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall
(l0 : L)
(r : R),
xL l0 < xR r)
(fxL : foralll0 : L,
(fun_ : GenNo S => A) (xL l0))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l0 : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l0)
(xR r)
(xcut l0 r)
(fxL l0)
(fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L' R' s'
yL yR ycut fyL fyR fycut) q: forallr : R', {{ xL | xR//xcut }} < yR r dq: forallr : R',
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) {{ xL | xR//xcut }}
(yR r) (q r)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall
(l : L)
(r0 : R),
xL l < xR r0)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr0 : R,
(fun_ : GenNo S => A) (xR r0))
(fxcut : forall
(l : L)
(r0 : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r0)
(xcut l r0)
(fxL l)
(fxR r0)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L R s xL
xR xcut fxL fxR fxcut)
(fyR r)
(fun (xy : GenNo S) (_ : x <= y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) => dle a b)
{{ xL | xR//xcut }} {{ yL | yR//ycut }}
(le_lr xL xR xcut yL yR ycut p q)
((fun (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall (l : L) (r : R),
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r) (xcut l r)
(fxL l) (fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L R s xL xR
xcut fxL fxR fxcut)
((fun (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall (l : L) (r : R),
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r) (xcut l r)
(fxL l) (fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L' R' s' yL
yR ycut fyL fyR fycut)
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) p: foralll : L, xL l < {{ yL | yR//ycut }} dp: foralll : L,
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) {{ yL | yR//ycut }}
(p l) (fxL l)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall
(l0 : L)
(r : R),
xL l0 < xR r)
(fxL : foralll0 : L,
(fun_ : GenNo S => A) (xL l0))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l0 : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l0)
(xR r)
(xcut l0 r)
(fxL l0)
(fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L' R' s'
yL yR ycut fyL fyR fycut) q: forallr : R', {{ xL | xR//xcut }} < yR r dq: forallr : R',
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) {{ xL | xR//xcut }}
(yR r) (q r)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall
(l : L)
(r0 : R),
xL l < xR r0)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr0 : R,
(fun_ : GenNo S => A) (xR r0))
(fxcut : forall
(l : L)
(r0 : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r0)
(xcut l r0)
(fxL l)
(fxR r0)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L R s xL
xR xcut fxL fxR fxcut)
(fyR r)
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)
apply dle_lr; assumption.
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) l: L' p: {{ xL | xR//xcut }} <= yL l dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) {{ xL | xR//xcut }}
(yL l) p
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L R s xL
xR xcut fxL fxR fxcut)
(fyL l)
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) => dlt a b)
{{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_l xL xR xcut yL yR ycut l p)
((fun (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall (l : L) (r : R),
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r) (xcut l r)
(fxL l) (fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L R s xL xR
xcut fxL fxR fxcut)
((fun (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall (l : L) (r : R),
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r) (xcut l r)
(fxL l) (fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L' R' s' yL
yR ycut fyL fyR fycut)
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) l: L' p: {{ xL | xR//xcut }} <= yL l dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) {{ xL | xR//xcut }}
(yL l) p
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L R s xL
xR xcut fxL fxR fxcut)
(fyL l)
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)
apply dlt_l with l; assumption.
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) r: R p: xR r <= {{ yL | yR//ycut }} dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) (xR r) {{ yL | yR//ycut }} p
(fxR r)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L' R' s'
yL yR ycut fyL fyR fycut)
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) => dlt a b)
{{ xL | xR//xcut }} {{ yL | yR//ycut }}
(lt_r xL xR xcut yL yR ycut r p)
((fun (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall (l : L) (r : R),
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r) (xcut l r)
(fxL l) (fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L R s xL xR
xcut fxL fxR fxcut)
((fun (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall (l : L) (r : R),
(fun (xy : GenNo S) (_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r) (xcut l r)
(fxL l) (fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L' R' s' yL
yR ycut fyL fyR fycut)
S: OptionSort No:= GenNo S: Type H: Funext A: Type dle: A -> A -> Type is_mere_relation0: is_mere_relation A dle dlt: A -> A -> Type is_mere_relation1: is_mere_relation A dlt dcut: forallLR : Type,
InSort S L R ->
forall (xL : L -> No)
(xR : R -> No),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A)
(fxR : R -> A),
(forall (l : L) (r : R), dlt (fxL l) (fxR r)) ->
A dpath: forallab : A, dle a b -> dle b a -> a = b dle_lr: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
dlt (fxL l)
(dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_l: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
dle (dcut L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) dlt_r: forall (LR : Type)
(s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L)
(r : R),
dlt (fxL l) (fxR r))
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L')
(r : R'),
dlt (fyL l) (fyR r))
(r : R),
xR r <= {{ yL | yR//ycut }} ->
dle (fxR r)
(dcut L' R' s' yL yR ycut fyL fyR fycut) ->
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut) L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: foralll : L, (fun_ : GenNo S => A) (xL l) fxR: forallr : R, (fun_ : GenNo S => A) (xR r) fxcut: forall (l : L) (r : R),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (xL l) (xR r)
(xcut l r) (fxL l)
(fxR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r fyL: foralll : L', (fun_ : GenNo S => A) (yL l) fyR: forallr : R', (fun_ : GenNo S => A) (yR r) fycut: forall (l : L') (r : R'),
(fun (xy : GenNo S)
(_ : x < y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dlt a b) (yL l) (yR r)
(ycut l r) (fyL l)
(fyR r) r: R p: xR r <= {{ yL | yR//ycut }} dp: (fun (xy : GenNo S)
(_ : x <= y) (a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y) =>
dle a b) (xR r) {{ yL | yR//ycut }} p
(fxR r)
((fun (LR : Type)
(s : InSort S L R)
(xL : L -> GenNo S)
(xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : foralll : L,
(fun_ : GenNo S => A) (xL l))
(fxR : forallr : R,
(fun_ : GenNo S => A) (xR r))
(fxcut : forall
(l : L)
(r : R),
(fun
(xy : GenNo S)
(_ : x < y)
(a : (fun_ : GenNo S => A) x)
(b : (fun_ : GenNo S => A) y)
=>
dlt a b)
(xL l)
(xR r)
(xcut l r)
(fxL l)
(fxR r)) =>
dcut L R s xL xR xcut fxL fxR fxcut) L' R' s'
yL yR ycut fyL fyR fycut)
dlt (dcut L R s xL xR xcut fxL fxR fxcut)
(dcut L' R' s' yL yR ycut fyL fyR fycut)
apply dlt_r with r; assumption.Defined.DefinitionNo_rec_le (xy : No) (p : x <= y)
: dle (No_rec x) (No_rec y)
:= No_ind_le (fun_ => A) (fun___ab => dle a b)
(fun___ab => dlt a b) _ _ _ _ _ x y p.DefinitionNo_rec_lt (xy : No) (p : x < y)
: dlt (No_rec x) (No_rec y)
:= No_ind_lt (fun_ => A) (fun___ab => dle a b)
(fun___ab => dlt a b) _ _ _ _ _ x y p.DefinitionNo_rec_package
: { f : No -> A &
(forall (xy : No) (p : x <= y), dle (f x) (f y)) *
(forall (xy : No) (p : x < y), dlt (f x) (f y)) }
:= ( No_rec ; (No_rec_le , No_rec_lt) ).DefinitionNo_rec_cut
(LR : Type@{i}) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
: No_rec {{ xL | xR // xcut }}
= dcut L R _ xL xR xcut
(funl => No_rec (xL l)) (funr => No_rec (xR r))
(funlr => No_rec_lt (xL l) (xR r) (xcut l r))
:= 1.EndNoRec.(** ** Conway's Theorem 0 *)(** First we prove that *if* a left option of [y] is [<=] itself, then it is [< y]. *)
S: OptionSort No:= GenNo S: Type H: Funext x: No xle: x <= x L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r l: L' p: x = yL l
x < {{ yL | yR//ycut }}
S: OptionSort No:= GenNo S: Type H: Funext x: No xle: x <= x L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r l: L' p: x = yL l
x < {{ yL | yR//ycut }}
S: OptionSort No:= GenNo S: Type H: Funext L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r l: L' L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll0 : L,
xL l0 <= xL l0 ->
xL l0 = yL l -> xL l0 < {{ yL | yR//ycut }} IHR: forallr : R,
xR r <= xR r ->
xR r = yL l -> xR r < {{ yL | yR//ycut }} xle: {{ xL | xR//xcut }} <= {{ xL | xR//xcut }} p: {{ xL | xR//xcut }} = yL l
{{ xL | xR//xcut }} < {{ yL | yR//ycut }}
S: OptionSort No:= GenNo S: Type H: Funext L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r l: L' L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll0 : L,
xL l0 <= xL l0 ->
xL l0 = yL l -> xL l0 < {{ yL | yR//ycut }} IHR: forallr : R,
xR r <= xR r ->
xR r = yL l -> xR r < {{ yL | yR//ycut }} xle: {{ xL | xR//xcut }} <= {{ xL | xR//xcut }} p: {{ xL | xR//xcut }} = yL l
{{ xL | xR//xcut }} <= yL l
exact (transport (funz => {{ xL | xR // xcut}} <= z) p xle).Defined.(** And dually *)
S: OptionSort No:= GenNo S: Type H: Funext x: No xle: x <= x L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r r: R' p: x = yR r
{{ yL | yR//ycut }} < x
S: OptionSort No:= GenNo S: Type H: Funext x: No xle: x <= x L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r r: R' p: x = yR r
{{ yL | yR//ycut }} < x
S: OptionSort No:= GenNo S: Type H: Funext L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r r: R' L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L,
xL l <= xL l ->
xL l = yR r -> {{ yL | yR//ycut }} < xL l IHR: forallr0 : R,
xR r0 <= xR r0 ->
xR r0 = yR r -> {{ yL | yR//ycut }} < xR r0 xle: {{ xL | xR//xcut }} <= {{ xL | xR//xcut }} p: {{ xL | xR//xcut }} = yR r
{{ yL | yR//ycut }} < {{ xL | xR//xcut }}
S: OptionSort No:= GenNo S: Type H: Funext L', R': Type s': InSort S L' R' yL: L' -> No yR: R' -> No ycut: forall (l : L') (r : R'), yL l < yR r r: R' L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L,
xL l <= xL l ->
xL l = yR r -> {{ yL | yR//ycut }} < xL l IHR: forallr0 : R,
xR r0 <= xR r0 ->
xR r0 = yR r -> {{ yL | yR//ycut }} < xR r0 xle: {{ xL | xR//xcut }} <= {{ xL | xR//xcut }} p: {{ xL | xR//xcut }} = yR r
yR r <= {{ xL | xR//xcut }}
exact (transport (funz => z <= {{ xL | xR // xcut}}) p xle).Defined.(** Theorem 0 Part (i) *)
S: OptionSort No:= GenNo S: Type H: Funext x: No
x <= x
S: OptionSort No:= GenNo S: Type H: Funext x: No
x <= x
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, xL l <= xL l IHR: forallr : R, xR r <= xR r
{{ xL | xR//xcut }} <= {{ xL | xR//xcut }}
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, xL l <= xL l IHR: forallr : R, xR r <= xR r
foralll : L, xL l < {{ xL | xR//xcut }}
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, xL l <= xL l IHR: forallr : R, xR r <= xR r
forallr : R, {{ xL | xR//xcut }} < xR r
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, xL l <= xL l IHR: forallr : R, xR r <= xR r
foralll : L, xL l < {{ xL | xR//xcut }}
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, xL l <= xL l IHR: forallr : R, xR r <= xR r l: L
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, xL l <= xL l IHR: forallr : R, xR r <= xR r
forallr : R, {{ xL | xR//xcut }} < xR r
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, xL l <= xL l IHR: forallr : R, xR r <= xR r r: R
{{ xL | xR//xcut }} < xR r
exact (Conway_theorem0_lemma2 (xR r) (IHR r) _ _ _ _ 1).Defined.#[export] Instancereflexive_le `{Funext} : Reflexive le
:= le_reflexive.(** Theorem 0 Part (ii), left half *)
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r l: L
xL l < {{ xL | xR//xcut }}
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r l: L
xL l < {{ xL | xR//xcut }}
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r l: L
xL l <= xL l
apply le_reflexive.Defined.(** Theorem 0 Part (ii), right half *)
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r r: R
{{ xL | xR//xcut }} < xR r
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r r: R
{{ xL | xR//xcut }} < xR r
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r r: R
xR r <= xR r
apply le_reflexive.Defined.
S: OptionSort No:= GenNo S: Type H: Funext
IsHSet No
S: OptionSort No:= GenNo S: Type H: Funext
IsHSet No
S: OptionSort No:= GenNo S: Type H: Funext
Reflexive (funxy : No => (x <= y) * (y <= x))
S: OptionSort No:= GenNo S: Type H: Funext
forallxy : No, (x <= y) * (y <= x) -> x = y
S: OptionSort No:= GenNo S: Type H: Funext
Reflexive (funxy : No => (x <= y) * (y <= x))
intros x; split; apply le_reflexive.
S: OptionSort No:= GenNo S: Type H: Funext
forallxy : No, (x <= y) * (y <= x) -> x = y
intros x y [xley ylex]; apply path_No; assumption.Defined.(** ** "One-line proofs" *)(** In particular, the proofs of cut-ness don't impact equality of surreals. However, in practice we generally need more than this: we need to be able to modify the indexing types along equivalences. *)
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No L', R': Type s': InSort S L' R' xL': L' -> No xR': R' -> No eL: L <~> L' eR: R <~> R' xLeq: foralll : L, xL l = xL' (eL l) xReq: forallr : R, xR r = xR' (eR r) xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L') (r : R'), xL' l < xR' r
{{ xL | xR//xcut }} = {{ xL' | xR'//xcut' }}
S: OptionSort No:= GenNo S: Type H: Funext L, R: Type s: InSort S L R xL: L -> No xR: R -> No L', R': Type s': InSort S L' R' xL': L' -> No xR': R' -> No eL: L <~> L' eR: R <~> R' xLeq: foralll : L, xL l = xL' (eL l) xReq: forallr : R, xR r = xR' (eR r) xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L') (r : R'), xL' l < xR' r
{{ xL | xR//xcut }} = {{ xL' | xR'//xcut' }}
apply path_No; apply le_lr;
[ intros l; rewrite xLeq
| intros r; rewrite <- (eisretr eR r), <- xReq
| intros l; rewrite <- (eisretr eL l), <- xLeq
| intros r; rewrite xReq ];
tryapply lt_lopt;
tryapply lt_ropt.Qed.Definitionpath_No_easy' `{Funext}
{L R : Type} {s : InSort S L R}
(xL xL' : L -> No) (xR xR' : R -> No)
(xLeq : foralll, xL l = xL' l)
(xReq : forallr, xR r = xR' r)
(xcut : forall (l:L) (r:R), (xL l) < (xR r))
: {{ xL | xR // xcut }}
= {{ xL' | xR' //
(funlr => transport (funxy => fst xy < snd xy)
(path_prod' (xLeq l) (xReq r))
(xcut l r)) }}
:= path_No_easy xL xR xL' xR' 11 xLeq xReq xcut _.EndOptionSort.(** When we want to do induction on several variables at once, we have to be careful to do them in the right order. This tactic does that, by calling itself recursively (although it doesn't choose useful names for all the hypotheses it introduces). We have to define this here outside of all sections so that it will be visible globally. *)Ltacrepeat_No_ind_hprop :=
trymatch goal with
| [ x : GenNo ?S |- _ ] =>
revert x;
repeat_No_ind_hprop;
refine (No_ind_hprop _ _);
intros ? ? ? ? ? ? ? ?
end.(** ** Encode-decode to characterize [<] and [<=] recursively (Theorem 11.6.7). *)SectionNoCodes.Context `{Univalence}.Universei.Context {S : OptionSort@{i}}.LetNo := GenNo S.LetA := {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)} }.SectionInner.Context {LR : Type@{i} } {s : InSort S L R}
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(xL_let : L -> A) (xR_let : R -> A)
(x_lt_le : forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y).LetA' (y : No) : Type
:= { lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y) }.LetA'le (yz : No) (p : y <= z) (tr : A' y) (sq : A' z) : Type
:= (fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1).LetA'lt (yz : No) (p : y < z) (tr : A' y) (sq : A' z) : Type
:= (fst tr.1 -> snd sq.1).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
{inner : forally : No, A' y &
(forall (yz : No) (p : y <= z),
A'le y z p (inner y) (inner z)) *
(forall (yz : No) (p : y < z),
A'lt y z p (inner y) (inner z))}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
{inner : forally : No, A' y &
(forall (yz : No) (p : y <= z),
A'le y z p (inner y) (inner z)) *
(forall (yz : No) (p : y < z),
A'lt y z p (inner y) (inner z))}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
abstract (
refine ((_,_),_);
[ intros h; strip_truncations;
destruct h as [[l' h]|[r h]]; split; intros;
[ refine (snd (fst (xL_let l).2.2) (yL l') y _ _);
[ refine (fst (fst (fst (xL_let l).2.2)) (yL l') _);
exact (snd (fst (x_let_yL l').2) l h)
| by (apply lt_lopt; exact _) ]
| exact (y_lt_le l' r' h)
| exact (x_lt_le l r y h)
| refine (snd (x_let_yR r').2 r _);
refine (fst (fst (fst (xR_let r).2.2)) _ _);
refine (snd (fst (xR_let r).2.2) y (yR r') h _);
apply lt_ropt; exact _ ]
| intros l [h k]; apply h
| intros r h; apply tr, inr; exact (r;h) ] ).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
abstract (
cbn;
intros L' R' ? yL yR ycut x_let_yL x_let_yR y_lt_le;
set (y := {{ yL | yR // ycut }});
intros L'' R'' ? zL zR zcut x_let_zL x_let_zR z_lt_le;
set (z := {{ zL | zR // zcut }});
intros yL_lt_z h1 y_lt_zR h2;
assert (y_le_z := le_lr yL yR ycut zL zR zcut yL_lt_z y_lt_zR);
split; [ intros [h3 h4]; split
| intros h3; strip_truncations;
destruct h3 as [[l' h3]|[r h3]] ] ;
[ intros l; exact (snd (xL_let l).2.2 y z (h3 l) y_le_z)
| intros r''; exact (h2 r'' (h3 , h4))
| exact (h1 l' h3)
| apply tr, inr; existsr;
exact (snd (fst (fst (xR_let r).2.2)) y z h3 y_le_z) ] ).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
abstract (
cbn;
intros L' R' ? yL yR ycut x_let_yL x_let_yR y_lt_le;
set (y := {{ yL | yR // ycut }});
intros L'' R'' ? zL zR zcut x_let_zL x_let_zR z_lt_le;
set (z := {{ zL | zR // zcut }});
intros l'' y_le_zL [h1 h2] x_le_y;
apply tr, inl; exact (l''; h1 x_le_y) ).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> No xR: R -> No xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R) (y : No),
(xR_let r).1 y -> ((xL_let l).2).1 y A':= funy : No =>
{lelt'_x_y : HProp * HProp &
(snd lelt'_x_y -> fst lelt'_x_y) *
(foralll : L, fst lelt'_x_y -> ((xL_let l).2).1 y) *
(forallr : R, (xR_let r).1 y -> snd lelt'_x_y)}: No -> Type A'le:= fun (yz : No) (_ : y <= z)
(tr : A' y) (sq : A' z) =>
(fst tr.1 -> fst sq.1) * (snd tr.1 -> snd sq.1): forallyz : No, y <= z -> A' y -> A' z -> Type A'lt:= fun (yz : No) (_ : y < z)
(tr : A' y) (sq : A' z) =>
fst tr.1 -> snd sq.1: forallyz : No, y < z -> A' y -> A' z -> Type
abstract (
cbn;
intros L' R' ? yL yR ycut x_let_yL x_let_yR y_lt_le;
set (y := {{ yL | yR // ycut }});
intros L'' R'' ? zL zR zcut x_let_zL x_let_zR z_lt_le;
set (z := {{ zL | zR // zcut }});
intros r' yR_le_z [h1 h2] x_le_y;
apply h2; exact (snd x_le_y r') ).Defined.Local Definitioninner (y : No) : A' y
:= inner_package.1 y.(** These computation laws hold definitionally, but it helps Coq out if we prove them explicitly and then rewrite along them later. *)Definitioninner_cut_le
(L'R' : Type@{i}) {s : InSort S L' R'}
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
: fst (inner {{ yL | yR // ycut }}).1 =
(Build_HProp ((foralll, (xL_let l).2.1 {{ yL | yR // ycut }}) *
(forallr', snd (inner (yR r')).1)))
:= 1.Definitioninner_cut_lt
(L'R' : Type@{i}) {s : InSort S L' R'}
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l:L') (r:R'), (yL l) < (yR r))
: snd (inner {{ yL | yR // ycut }}).1 =
(hor {l':L' & fst (inner (yL l')).1}
{r:R & (xR_let r).1 {{ yL | yR // ycut }} })
:= 1.Local Definitioninner_leyzp : A'le y z p (inner y) (inner z)
:= fst (inner_package.2) y z p.Local Definitioninner_ltyzp : A'lt y z p (inner y) (inner z)
:= snd (inner_package.2) y z p.EndInner.(** We instruct [simpl]/[cbn] not to unfold [inner]. We will do the "unfolding" ourselves by rewriting along [inner_cut_le] and [inner_cut_lt], so as to keep better control over the resulting terms (and particularly their size). *)Arguments inner : simpl never.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
{lelt : No -> A &
(forallxy : No,
x <= y ->
forallz : No,
((lelt y).1 z -> (lelt x).1 z) *
(((lelt y).2).1 z -> ((lelt x).2).1 z)) *
(forallxy : No,
x < y ->
forallz : No, (lelt y).1 z -> ((lelt x).2).1 z)}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
{lelt : No -> A &
(forallxy : No,
x <= y ->
forallz : No,
((lelt y).1 z -> (lelt x).1 z) *
(((lelt y).2).1 z -> ((lelt x).2).1 z)) *
(forallxy : No,
x < y ->
forallz : No, (lelt y).1 z -> ((lelt x).2).1 z)}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forallLR : Type,
InSort S L R ->
forall (xL : L -> GenNo S) (xR : R -> GenNo S),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A) (fxR : R -> A),
(forall (l : L) (r : R),
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
(fxL l) (fxR r)) -> A
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forallab : A,
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y)) a b ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y)) b a ->
a = b
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo S)
(yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L') (r : R'),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fyL l)
(fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
(fxL l) (?dcut L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
(?dcut L R s xL xR xcut fxL fxR fxcut) (fyR r)) ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y))
(?dcut L R s xL xR xcut fxL fxR fxcut)
(?dcut L' R' s' yL yR ycut fyL fyR fycut)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo S)
(yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L') (r : R'),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fyL l)
(fyR r)) (l : L'),
{{ xL | xR//xcut }} <= yL l ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y))
(?dcut L R s xL xR xcut fxL fxR fxcut) (fyL l) ->
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
(?dcut L R s xL xR xcut fxL fxR fxcut)
(?dcut L' R' s' yL yR ycut fyL fyR fycut)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo S)
(yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L') (r : R'),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fyL l)
(fyR r)) (r : R),
xR r <= {{ yL | yR//ycut }} ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y))
(fxR r) (?dcut L' R' s' yL yR ycut fyL fyR fycut) ->
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
(?dcut L R s xL xR xcut fxL fxR fxcut)
(?dcut L' R' s' yL yR ycut fyL fyR fycut)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forallLR : Type,
InSort S L R ->
forall (xL : L -> GenNo S) (xR : R -> GenNo S),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> A) (fxR : R -> A),
(forall (l : L) (r : R),
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
(fxL l) (fxR r)) -> A
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r)
A
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r) x:= {{ xL | xR//xcut }}: GenNo S
A
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r) x:= {{ xL | xR//xcut }}: GenNo S
{lt'_x : No -> HProp &
(forally : No,
lt'_x y -> fst (inner xL_let xR_let x_lt_le y).1) *
(forallyz : No,
fst (inner xL_let xR_let x_lt_le y).1 ->
y <= z -> fst (inner xL_let xR_let x_lt_le z).1) *
(forallyz : No,
fst (inner xL_let xR_let x_lt_le y).1 ->
y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xL_let: L -> A xR_let: R -> A x_lt_le: forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r) x:= {{ xL | xR//xcut }}: GenNo S
(forally : No,
snd (inner xL_let xR_let x_lt_le y).1 ->
fst (inner xL_let xR_let x_lt_le y).1) *
(forallyz : No,
fst (inner xL_let xR_let x_lt_le y).1 ->
y <= z -> fst (inner xL_let xR_let x_lt_le z).1) *
(forallyz : No,
fst (inner xL_let xR_let x_lt_le y).1 ->
y < z -> snd (inner xL_let xR_let x_lt_le z).1) *
(forallyz : No,
snd (inner xL_let xR_let x_lt_le y).1 ->
y <= z -> snd (inner xL_let xR_let x_lt_le z).1)
abstract (
repeatsplit;
[ intros y; exact (fst (fst (inner xL_let xR_let x_lt_le y).2))
| intros y z x_le_y y_le_z;
exact (fst (inner_le xL_let xR_let x_lt_le y z y_le_z) x_le_y)
| intros y z x_le_y y_lt_z;
exact (inner_lt xL_let xR_let x_lt_le y z y_lt_z x_le_y)
| intros y z x_lt_y y_le_z;
exact (snd (inner_le xL_let xR_let x_lt_le y z y_le_z) x_lt_y) ]).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forallab : A,
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y)) a b ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y)) b a ->
a = b
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo S)
(yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L') (r : R'),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fyL l)
(fyR r)),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
(fxL l)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l0 : L0) (r : R0),
xL0 l0 < xR0 r) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l0 : L0) (r : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l0) (xR_let r)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L' R' s' yL yR ycut fyL fyR fycut)) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l : L0) (r0 : R0),
xL0 l < xR0 r0) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l : L0) (r0 : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r0)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L R s xL xR xcut fxL fxR fxcut)
(fyR r)) ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y))
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l : L0) (r : R0), xL0 l < xR0 r)
(xL_let : L0 -> A) (xR_let : R0 -> A)
(x_lt_le : forall (l : L0) (r : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L R s xL xR xcut fxL fxR fxcut)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l : L0) (r : R0), xL0 l < xR0 r)
(xL_let : L0 -> A) (xR_let : R0 -> A)
(x_lt_le : forall (l : L0) (r : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L' R' s' yL yR ycut fyL fyR fycut)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo S)
(yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L') (r : R'),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fyL l)
(fyR r)) (l : L'),
{{ xL | xR//xcut }} <= yL l ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y))
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l0 : L0) (r : R0),
xL0 l0 < xR0 r) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l0 : L0) (r : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l0) (xR_let r)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L R s xL xR xcut fxL fxR fxcut)
(fyL l) ->
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l0 : L0) (r : R0),
xL0 l0 < xR0 r) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l0 : L0) (r : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l0) (xR_let r)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L R s xL xR xcut fxL fxR fxcut)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l0 : L0) (r : R0),
xL0 l0 < xR0 r) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l0 : L0) (r : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l0) (xR_let r)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L' R' s' yL yR ycut fyL fyR fycut)
abstract (
intros L R ? xL xR xcut xL_let xR_let x_le_lt
L' R' ? yL yR ycut yL_let yR_let y_le_lt;
set (x := {{ xL | xR // xcut }});
set (y := {{ yL | yR // ycut }});
cbn; intros l x_le_yL zH;
refine (No_ind_hprop _ _);
intros L'' R'' ? zL zR zcut zLH zRH y_le_z;
refine (snd (zH {{ zL | zR // zcut }}) _);
rewrite inner_cut_le in y_le_z;
exact (fst y_le_z l) ).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> A) (fxR : R -> A)
(fxcut : forall (l : L) (r : R),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fxL l)
(fxR r)) (L'R' : Type)
(s' : InSort S L' R') (yL : L' -> GenNo S)
(yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> A) (fyR : R' -> A)
(fycut : forall (l : L') (r : R'),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y) (fyL l)
(fyR r)) (r : R),
xR r <= {{ yL | yR//ycut }} ->
(fundmht : A =>
forally : No,
(ht.1 y -> dm.1 y) * ((ht.2).1 y -> (dm.2).1 y))
(fxR r)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l : L0) (r0 : R0),
xL0 l < xR0 r0) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l : L0) (r0 : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r0)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L' R' s' yL yR ycut fyL fyR fycut) ->
(fundmht : A => forally : No, ht.1 y -> (dm.2).1 y)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l : L0) (r0 : R0),
xL0 l < xR0 r0) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l : L0) (r0 : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r0)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L R s xL xR xcut fxL fxR fxcut)
((fun (L0R0 : Type) (s0 : InSort S L0 R0)
(xL0 : L0 -> GenNo S) (xR0 : R0 -> GenNo S)
(xcut0 : forall (l : L0) (r0 : R0),
xL0 l < xR0 r0) (xL_let : L0 -> A)
(xR_let : R0 -> A)
(x_lt_le : forall (l : L0) (r0 : R0),
(fundmht : A =>
forally : No, ht.1 y -> (dm.2).1 y)
(xL_let l) (xR_let r0)) =>
letx := {{ xL0 | xR0//xcut0 }} in
(funy : No =>
fst (inner xL_let xR_let x_lt_le y).1;
funy : No =>
snd (inner xL_let xR_let x_lt_le y).1;
No_codes_package_subproof L0 R0 xL_let xR_let
x_lt_le)) L' R' s' yL yR ycut fyL fyR fycut)
abstract (
intros L R ? xL xR xcut xL_let xR_let x_le_lt
L' R' ? yL yR ycut yL_let yR_let y_le_lt;
set (x := {{ xL | xR // xcut }});
set (y := {{ yL | yR // ycut }});
cbn; intros r xR_le_y zH;
refine (No_ind_hprop _ _);
intros L'' R'' ? zL zR zcut zLH zRH y_le_z;
rewrite inner_cut_lt;
apply tr; right; existsr;
refine (fst (zH {{ zL | zR // zcut }}) y_le_z) ).Defined.Definitionle' (xy : No) : HProp
:= (No_codes_package.1 x).1 y.Definitionlt' (xy : No) : HProp
:= (No_codes_package.1 x).2.1 y.Definitionlt'_le'xy
: lt' x y -> le' x y
:= (fst (fst (fst (No_codes_package.1 x).2.2)) y).Definitionle_le'_transxyz
: x <= y -> le' y z -> le' x z
:= funpq => (fst (fst (No_codes_package.2) x y p z) q).Definitionle_lt'_transxyz
: x <= y -> lt' y z -> lt' x z
:= funpq => (snd (fst (No_codes_package.2) x y p z) q).Definitionlt_le'_transxyz
: x < y -> le' y z -> lt' x z
:= funpq => (snd (No_codes_package.2) x y p z q).Definitionle'_le_transxyz
: le' x y -> y <= z -> le' x z
:= funpq => (snd (fst (fst (No_codes_package.1 x).2.2)) y z p q).Definitionle'_lt_transxyz
: le' x y -> y < z -> lt' x z
:= funpq => (snd (fst (No_codes_package.1 x).2.2) y z p q).Definitionlt'_le_transxyz
: lt' x y -> y <= z -> lt' x z
:= funpq => (snd (No_codes_package.1 x).2.2 y z p q).(** These computation laws hold definitionally, but it takes Coq a little while to verify that. Thus, we prove them once and then [rewrite] along them later, so we don't have to do the verification every time. *)Definitionle'_cut
(LR : Type) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
: le' {{xL | xR // xcut}} {{yL | yR // ycut}}
= ((foralll, lt' (xL l) {{ yL | yR // ycut }}) *
(forallr', lt' {{ xL | xR // xcut }} (yR r')))
(** For some reason, Coq has a really hard time checking the version of this that asserts an equality in [HProp]. But fortunately, we only ever really need the equality of types. *)
:> Type
:= 1.Definitionlt'_cut
(LR : Type) (s : InSort S L R)
(xL : L -> No) (xR : R -> No)
(xcut : forall (l : L) (r : R), xL l < xR r)
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> No) (yR : R' -> No)
(ycut : forall (l : L') (r : R'), yL l < yR r)
: lt' {{xL | xR // xcut}} {{yL | yR // ycut}}
= (hor {l':L' & le' {{ xL | xR // xcut }} (yL l')}
{r:R & le' (xR r) {{ yL | yR // ycut }} })
:= 1.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
(x <= y -> le' x y) * (x < y -> lt' x y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
(x <= y -> le' x y) * (x < y -> lt' x y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
forallLR : Type,
InSort S L R ->
forall (xL : L -> GenNo S) (xR : R -> GenNo S),
(forall (l : L) (r : R), xL l < xR r) ->
(L -> Unit) ->
(R -> Unit) ->
(forall (l : L) (r : R), lt' (xL l) (xR r)) -> Unit
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
forall (xy : GenNo S) (ab : Unit) (p : x <= y)
(q : y <= x),
le' x y ->
le' y x ->
transport (fun_ : GenNo S => Unit) (path_No x y p q)
a = b
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
forallLR : Type,
InSort S L R ->
forall (xL : L -> GenNo S) (xR : R -> GenNo S),
(forall (l : L) (r : R), xL l < xR r) ->
(L -> Unit) ->
(R -> Unit) ->
(forall (l : L) (r : R), lt' (xL l) (xR r)) -> Unit
intros; exact tt.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
forall (xy : GenNo S) (ab : Unit) (p : x <= y)
(q : y <= x),
le' x y ->
le' y x ->
transport (fun_ : GenNo S => Unit) (path_No x y p q)
a = b
intros; apply path_contr.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) l: L' x_le_yL: {{ xL | xR//xcut }} <= yL l x_le_yL': le' {{ xL | xR//xcut }} (yL l)
lt' {{ xL | xR//xcut }} {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) l: L' x_le_yL: {{ xL | xR//xcut }} <= yL l x_le_yL': le' {{ xL | xR//xcut }} (yL l)
hor {l' : L' & le' {{ xL | xR//xcut }} (yL l')}
{r : R & le' (xR r) {{ yL | yR//ycut }}}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) l: L' x_le_yL: {{ xL | xR//xcut }} <= yL l x_le_yL': le' {{ xL | xR//xcut }} (yL l)
{l' : L' & le' {{ xL | xR//xcut }} (yL l')}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) l: L' x_le_yL: {{ xL | xR//xcut }} <= yL l x_le_yL': le' {{ xL | xR//xcut }} (yL l)
le' {{ xL | xR//xcut }} (yL l)
exact x_le_yL'.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) r: R xR_le_y: xR r <= {{ yL | yR//ycut }} xR_le_y': le' (xR r) {{ yL | yR//ycut }}
lt' {{ xL | xR//xcut }} {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) r: R xR_le_y: xR r <= {{ yL | yR//ycut }} xR_le_y': le' (xR r) {{ yL | yR//ycut }}
hor {l' : L' & le' {{ xL | xR//xcut }} (yL l')}
{r : R & le' (xR r) {{ yL | yR//ycut }}}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) r: R xR_le_y: xR r <= {{ yL | yR//ycut }} xR_le_y': le' (xR r) {{ yL | yR//ycut }}
{r : R & le' (xR r) {{ yL | yR//ycut }}}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xcut': forall (l : L) (r : R), lt' (xL l) (xR r) L', R': Type s': InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r ycut': forall (l : L') (r : R'), lt' (yL l) (yR r) r: R xR_le_y: xR r <= {{ yL | yR//ycut }} xR_le_y': le' (xR r) {{ yL | yR//ycut }}
le' (xR r) {{ yL | yR//ycut }}
exact xR_le_y'.Qed.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
(le' x y -> x <= y) * (lt' x y -> x < y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No
(le' x y -> x <= y) * (lt' x y -> x < y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type
forallxy : No,
(le' x y -> x <= y) * (lt' x y -> x < y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y)
(* Coq can find this proof by typeclass search, but helping it out makes this faster. *)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) x_lt_y: lt' {{ xL | xR//xcut }} {{ yL | yR//ycut }}
{{ xL | xR//xcut }} < {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) x_lt_y: hor
{l' : L' & le' {{ xL | xR//xcut }} (yL l')}
{r : R & le' (xR r) {{ yL | yR//ycut }}}
{{ xL | xR//xcut }} < {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) l: L' x_le_yL: le' {{ xL | xR//xcut }} (yL l)
{{ xL | xR//xcut }} < {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) r: R xR_le_y: le' (xR r) {{ yL | yR//ycut }}
{{ xL | xR//xcut }} < {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) l: L' x_le_yL: le' {{ xL | xR//xcut }} (yL l)
{{ xL | xR//xcut }} < {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) l: L' x_le_yL: le' {{ xL | xR//xcut }} (yL l)
{{ xL | xR//xcut }} <= yL l
exact (fst (yHL l) x_le_yL).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) r: R xR_le_y: le' (xR r) {{ yL | yR//ycut }}
{{ xL | xR//xcut }} < {{ yL | yR//ycut }}
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r xHL: forall (l : L) (y : No),
(le' (xL l) y -> xL l <= y) *
(lt' (xL l) y -> xL l < y) xHR: forall (r : R) (y : No),
(le' (xR r) y -> xR r <= y) *
(lt' (xR r) y -> xR r < y) L', R': Type s0: InSort S L' R' yL: L' -> GenNo S yR: R' -> GenNo S ycut: forall (l : L') (r : R'), yL l < yR r yHL: foralll : L',
(le' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} <= yL l) *
(lt' {{ xL | xR//xcut }} (yL l) ->
{{ xL | xR//xcut }} < yL l) yHR: forallr : R',
(le' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} <= yR r) *
(lt' {{ xL | xR//xcut }} (yR r) ->
{{ xL | xR//xcut }} < yR r) r: R xR_le_y: le' (xR r) {{ yL | yR//ycut }}
xR r <= {{ yL | yR//ycut }}
exact (fst (xHR r _) xR_le_y).Qed.DefinitionNo_encode_lexy := fst (No_encode_le_lt x y).DefinitionNo_encode_ltxy := snd (No_encode_le_lt x y).DefinitionNo_decode_lexy := fst (No_decode_le_lt x y).DefinitionNo_decode_ltxy := snd (No_decode_le_lt x y).
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No p: x < y
x <= y
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No p: x < y
x <= y
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No p: x < y
le' x y
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No p: x < y
lt' x y
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y: No p: x < y
x < y
assumption.Qed.(** Conway's theorem 1 *)
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No
x <= y -> y <= z -> x <= z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No
x <= y -> y <= z -> x <= z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y <= z
x <= z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y <= z
le' x z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y <= z
le' y z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y <= z
y <= z
assumption.Qed.#[export] Instancetrans_le : Transitive le
:= @le_le_trans.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No
x <= y -> y < z -> x < z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No
x <= y -> y < z -> x < z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y < z
x < z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y < z
lt' x z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y < z
lt' y z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x <= y q: y < z
y < z
assumption.Qed.
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No
x < y -> y <= z -> x < z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No
x < y -> y <= z -> x < z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x < y q: y <= z
x < z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x < y q: y <= z
lt' x z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x < y q: y <= z
le' y z
H: Univalence S: OptionSort No:= GenNo S: Type A:= {le'_x : No -> HProp &
{lt'_x : No -> HProp &
(forally : No, lt'_x y -> le'_x y) *
(forallyz : No, le'_x y -> y <= z -> le'_x z) *
(forallyz : No, le'_x y -> y < z -> lt'_x z) *
(forallyz : No, lt'_x y -> y <= z -> lt'_x z)}}: Type x, y, z: No p: x < y q: y <= z
y <= z
assumption.Qed.Definitionlt_lt_trans {xyz : No}
: (x < y) -> (y < z) -> (x < z)
:= funpq => lt_le_trans p (lt_le q).#[export] Instancetrans_lt : Transitive lt
:= @lt_lt_trans.EndNoCodes.(** ** Changing option sorts *)(** There is of course a "maximal" option sort, which defines "the" surreal numbers as in the book. *)DefinitionMaxSort : OptionSort := fun__ => Unit.DefinitionNo : Type := GenNo MaxSort.(** This instance should be the one found by default, so that cuts live in [No] unless otherwise specified. Thus, all other global instances of [InSort] should be declared with higher priority. *)Instanceinsort_maxsort {LR : Type}
: InSort MaxSort L R | 0
:= tt.(** Furthermore, every other kind of surreal number *embeds* into the maximal ones. So the other kinds of surreal numbers are really just subsets of the usual set of surreal numbers; but I don't know of a good way to define them except as their own HIITs. *)SectionRaiseSort.Context `{ua: Univalence} `{S : OptionSort}.
ua: Univalence S: OptionSort
GenNo S -> No
ua: Univalence S: OptionSort
GenNo S -> No
ua: Univalence S: OptionSort
forallLR : Type,
InSort S L R ->
forall (xL : L -> GenNo S) (xR : R -> GenNo S),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> No) (fxR : R -> No),
(forall (l : L) (r : R), fxL l < fxR r) -> No
ua: Univalence S: OptionSort
forallab : No, a <= b -> b <= a -> a = b
ua: Univalence S: OptionSort
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> No) (fxR : R -> No)
(fxcut : forall (l : L) (r : R), fxL l < fxR r)
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo S) (yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> No) (fyR : R' -> No)
(fycut : forall (l : L') (r : R'), fyL l < fyR r),
(foralll : L, xL l < {{ yL | yR//ycut }}) ->
(foralll : L,
fxL l < ?dcut L' R' s' yL yR ycut fyL fyR fycut) ->
(forallr : R', {{ xL | xR//xcut }} < yR r) ->
(forallr : R',
?dcut L R s xL xR xcut fxL fxR fxcut < fyR r) ->
?dcut L R s xL xR xcut fxL fxR fxcut <=
?dcut L' R' s' yL yR ycut fyL fyR fycut
ua: Univalence S: OptionSort
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> No) (fxR : R -> No)
(fxcut : forall (l : L) (r : R), fxL l < fxR r)
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo S) (yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> No) (fyR : R' -> No)
(fycut : forall (l : L') (r : R'), fyL l < fyR r)
(l : L'),
{{ xL | xR//xcut }} <= yL l ->
?dcut L R s xL xR xcut fxL fxR fxcut <= fyL l ->
?dcut L R s xL xR xcut fxL fxR fxcut <
?dcut L' R' s' yL yR ycut fyL fyR fycut
ua: Univalence S: OptionSort
forall (LR : Type) (s : InSort S L R)
(xL : L -> GenNo S) (xR : R -> GenNo S)
(xcut : forall (l : L) (r : R), xL l < xR r)
(fxL : L -> No) (fxR : R -> No)
(fxcut : forall (l : L) (r : R), fxL l < fxR r)
(L'R' : Type) (s' : InSort S L' R')
(yL : L' -> GenNo S) (yR : R' -> GenNo S)
(ycut : forall (l : L') (r : R'), yL l < yR r)
(fyL : L' -> No) (fyR : R' -> No)
(fycut : forall (l : L') (r : R'), fyL l < fyR r)
(r : R),
xR r <= {{ yL | yR//ycut }} ->
fxR r <= ?dcut L' R' s' yL yR ycut fyL fyR fycut ->
?dcut L R s xL xR xcut fxL fxR fxcut <
?dcut L' R' s' yL yR ycut fyL fyR fycut
ua: Univalence S: OptionSort
forallLR : Type,
InSort S L R ->
forall (xL : L -> GenNo S) (xR : R -> GenNo S),
(forall (l : L) (r : R), xL l < xR r) ->
forall (fxL : L -> No) (fxR : R -> No),
(forall (l : L) (r : R), fxL l < fxR r) -> No
ua: Univalence S: OptionSort L, R: Type s: InSort S L R xL: L -> GenNo S xR: R -> GenNo S xcut: forall (l : L) (r : R), xL l < xR r fxL: L -> No fxR: R -> No fxcut: forall (l : L) (r : R), fxL l < fxR r
ua: Univalence S: OptionSort x, y: GenNo S e: No_raise x = No_raise y
x <= y
ua: Univalence S: OptionSort x, y: GenNo S e: No_raise x = No_raise y
y <= x
ua: Univalence S: OptionSort x, y: GenNo S e: No_raise x = No_raise y
x <= y
ua: Univalence S: OptionSort x, y: GenNo S e: No_raise x = No_raise y
No_raise x <= No_raise y
rewrite e; apply reflexive_le.
ua: Univalence S: OptionSort x, y: GenNo S e: No_raise x = No_raise y
y <= x
ua: Univalence S: OptionSort x, y: GenNo S e: No_raise x = No_raise y
No_raise y <= No_raise x
rewrite e; apply reflexive_le.Qed.EndRaiseSort.(** ** Ordinals *)(** The type of "plump ordinals" can be identified with surreal numbers that hereditarily have no right options. *)DefinitionOrdSort : OptionSort := funLR => ~R.DefinitionPOrd := GenNo OrdSort.Instanceinsort_ordsort {L : Type}
: InSort OrdSort L Empty | 100
:= idmap.(** ** Decidable options *)(** A particularly interesting option sort restricts [L] and [R] to be decidable, i.e. either inhabited or empty. *)DefinitionDecSort : OptionSort
:= funLR => Decidable L * Decidable R.DefinitionDecNo : Type := GenNo DecSort.Instanceinsort_decsort {LR : Type}
{dl : Decidable L} {dr : Decidable R}
: InSort DecSort L R | 100
:= (dl , dr).(** Perhaps surprisingly, this is not a restriction at all! Any surreal number can be presented by a cut in which all the option sorts are hereditarily decidable. The basic idea is that we can always add a "sufficiently large" right option and a "sufficiently small" left option in order to make both families of options inhabited without changing the value of the cut, but the details are a bit tricky. *)
H: Univalence L, R: Type s: InSort MaxSort L R xL: L -> GenNo MaxSort xR: R -> GenNo MaxSort xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, hfiber No_raise (xL l) IHR: forallr : R, hfiber No_raise (xR r) uLR:= Unit + (L + R): Type X: Decidable uLR xLR:= sum_ind (fun_ : Unit + (L + R) => DecNo)
(unit_name zero)
(sum_ind (fun_ : L + R => DecNo)
(funl : L => (IHL l).1)
(funr : R => (IHR r).1))
:
uLR -> DecNo: uLR -> DecNo z:= {{ xLR | Empty_rec//rempty_cut }}: GenNo DecSort z':= {{ unit_name z | Empty_rec//rempty_cut }}: GenNo DecSort y:= {{ Empty_rec | xLR//lempty_cut }}: GenNo DecSort y':= {{ Empty_rec | unit_name y//lempty_cut }}: GenNo DecSort L':= Unit + L: Type X0: Decidable L' wL:= sum_ind (fun_ : Unit + L => DecNo)
(unit_name y') (funl : L => (IHL l).1)
:
L' -> DecNo: L' -> DecNo R':= Unit + R: Type X1: Decidable R' wR:= sum_ind (fun_ : Unit + R => DecNo)
(unit_name z') (funr : R => (IHR r).1)
:
R' -> DecNo: R' -> DecNo wcut: forall (l : L') (r : R'), wL l < wR r rwcut: forall (l : L') (r : R'),
No_raise (wL l) < No_raise (wR r)
No_raise z < No_raise (wR (inl tt))
H: Univalence L, R: Type s: InSort MaxSort L R xL: L -> GenNo MaxSort xR: R -> GenNo MaxSort xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, hfiber No_raise (xL l) IHR: forallr : R, hfiber No_raise (xR r) uLR:= Unit + (L + R): Type X: Decidable uLR xLR:= sum_ind (fun_ : Unit + (L + R) => DecNo)
(unit_name zero)
(sum_ind (fun_ : L + R => DecNo)
(funl : L => (IHL l).1)
(funr : R => (IHR r).1))
:
uLR -> DecNo: uLR -> DecNo z:= {{ xLR | Empty_rec//rempty_cut }}: GenNo DecSort z':= {{ unit_name z | Empty_rec//rempty_cut }}: GenNo DecSort y:= {{ Empty_rec | xLR//lempty_cut }}: GenNo DecSort y':= {{ Empty_rec | unit_name y//lempty_cut }}: GenNo DecSort L':= Unit + L: Type X0: Decidable L' wL:= sum_ind (fun_ : Unit + L => DecNo)
(unit_name y') (funl : L => (IHL l).1)
:
L' -> DecNo: L' -> DecNo R':= Unit + R: Type X1: Decidable R' wR:= sum_ind (fun_ : Unit + R => DecNo)
(unit_name z') (funr : R => (IHR r).1)
:
R' -> DecNo: R' -> DecNo wcut: forall (l : L') (r : R'), wL l < wR r rwcut: forall (l : L') (r : R'),
No_raise (wL l) < No_raise (wR r)
z < wR (inl tt)
refine (lt_lopt _ _ _ tt).
H: Univalence L, R: Type s: InSort MaxSort L R xL: L -> GenNo MaxSort xR: R -> GenNo MaxSort xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, hfiber No_raise (xL l) IHR: forallr : R, hfiber No_raise (xR r) uLR:= Unit + (L + R): Type X: Decidable uLR xLR:= sum_ind (fun_ : Unit + (L + R) => DecNo)
(unit_name zero)
(sum_ind (fun_ : L + R => DecNo)
(funl : L => (IHL l).1)
(funr : R => (IHR r).1))
:
uLR -> DecNo: uLR -> DecNo z:= {{ xLR | Empty_rec//rempty_cut }}: GenNo DecSort z':= {{ unit_name z | Empty_rec//rempty_cut }}: GenNo DecSort y:= {{ Empty_rec | xLR//lempty_cut }}: GenNo DecSort y':= {{ Empty_rec | unit_name y//lempty_cut }}: GenNo DecSort L':= Unit + L: Type X0: Decidable L' wL:= sum_ind (fun_ : Unit + L => DecNo)
(unit_name y') (funl : L => (IHL l).1)
:
L' -> DecNo: L' -> DecNo R':= Unit + R: Type X1: Decidable R' wR:= sum_ind (fun_ : Unit + R => DecNo)
(unit_name z') (funr : R => (IHR r).1)
:
R' -> DecNo: R' -> DecNo wcut: forall (l : L') (r : R'), wL l < wR r rwcut: forall (l : L') (r : R'),
No_raise (wL l) < No_raise (wR r) r: R
{{ xL | xR//xcut }} < No_raise (wR (inr r))
H: Univalence L, R: Type s: InSort MaxSort L R xL: L -> GenNo MaxSort xR: R -> GenNo MaxSort xcut: forall (l : L) (r : R), xL l < xR r IHL: foralll : L, hfiber No_raise (xL l) IHR: forallr : R, hfiber No_raise (xR r) uLR:= Unit + (L + R): Type X: Decidable uLR xLR:= sum_ind (fun_ : Unit + (L + R) => DecNo)
(unit_name zero)
(sum_ind (fun_ : L + R => DecNo)
(funl : L => (IHL l).1)
(funr : R => (IHR r).1))
:
uLR -> DecNo: uLR -> DecNo z:= {{ xLR | Empty_rec//rempty_cut }}: GenNo DecSort z':= {{ unit_name z | Empty_rec//rempty_cut }}: GenNo DecSort y:= {{ Empty_rec | xLR//lempty_cut }}: GenNo DecSort y':= {{ Empty_rec | unit_name y//lempty_cut }}: GenNo DecSort L':= Unit + L: Type X0: Decidable L' wL:= sum_ind (fun_ : Unit + L => DecNo)
(unit_name y') (funl : L => (IHL l).1)
:
L' -> DecNo: L' -> DecNo R':= Unit + R: Type X1: Decidable R' wR:= sum_ind (fun_ : Unit + R => DecNo)
(unit_name z') (funr : R => (IHR r).1)
:
R' -> DecNo: R' -> DecNo wcut: forall (l : L') (r : R'), wL l < wR r rwcut: forall (l : L') (r : R'),
No_raise (wL l) < No_raise (wR r) r: R
{{ xL | xR//xcut }} < xR r
exact (lt_ropt _ _ _ r).Defined.Definitionequiv_DecNo_raise `{Univalence}
: DecNo <~> No
:= Build_Equiv _ _ No_raise _.(** Note that this does not extend to other sorts. For instance, it is *not* true that any plump ordinal is equal to a cut whose types of left and right options are respectively hereditarily decidable and hereditarily empty. In particular, when making the type of left options inhabited, we have to use surreals whose type of right options is also inhabited. For instance, [{{ fun _:P => zero | Empty_rec // rempty_cut }}], for a proposition [P], is a plump ordinal, but to make its left options inhabited we have to use a negative surreal, which is not itself a plump ordinal. *)