annotate src/zorn.agda @ 799:c8a166abcae0

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sun, 07 Aug 2022 18:39:18 +0900
parents 9cf74877efab
children 06eedb0d26a0
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
478
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 477
diff changeset
1 {-# OPTIONS --allow-unsolved-metas #-}
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
2 open import Level hiding ( suc ; zero )
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
3 open import Ordinals
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
4 open import Relation.Binary
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
5 open import Relation.Binary.Core
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
6 open import Relation.Binary.PropositionalEquality
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
7 import OD
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
8 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
9
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
10 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
11 -- Zorn-lemma : { A : HOD }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
12 -- → o∅ o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
13 -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
14 -- → Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
15 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
16
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
17 open import zf
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
18 open import logic
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
19 -- open import partfunc {n} O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
20
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
21 open import Relation.Nullary
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
22 open import Data.Empty
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
23 import BAlgbra
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
24
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
25 open import Data.Nat hiding ( _<_ ; _≤_ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
26 open import Data.Nat.Properties
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
27 open import nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
28
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
29
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
30 open inOrdinal O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
31 open OD O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
32 open OD.OD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 open ODAxiom odAxiom
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
34 import OrdUtil
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
35 import ODUtil
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
36 open Ordinals.Ordinals O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
37 open Ordinals.IsOrdinals isOrdinal
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
38 open Ordinals.IsNext isNext
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
39 open OrdUtil O
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
40 open ODUtil O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
41
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
42
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
43 import ODC
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
44
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
45 open _∧_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
46 open _∨_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
47 open Bool
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
48
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
49 open HOD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
50
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
51 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
52 -- Partial Order on HOD ( possibly limited in A )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
53 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
54
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
55 _<<_ : (x y : Ordinal ) → Set n -- Set n order
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
56 x << y = * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
57
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
58 _<=_ : (x y : Ordinal ) → Set n -- Set n order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
59 x <= y = (x ≡ y ) ∨ ( * x < * y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
60
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
61 POO : IsStrictPartialOrder _≡_ _<<_
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
62 POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
63 ; trans = IsStrictPartialOrder.trans PO
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
64 ; irrefl = λ x=y x<y → IsStrictPartialOrder.irrefl PO (cong (*) x=y) x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
65 ; <-resp-≈ = record { fst = λ {x} {y} {y1} y=y1 xy1 → subst (λ k → x << k ) y=y1 xy1 ; snd = λ {x} {x1} {y} x=x1 x1y → subst (λ k → k << x ) x=x1 x1y } }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
66
528
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
67 _≤_ : (x y : HOD) → Set (Level.suc n)
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
68 x ≤ y = ( x ≡ y ) ∨ ( x < y )
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
69
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
70 ≤-ftrans : {x y z : HOD} → x ≤ y → y ≤ z → x ≤ z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
71 ≤-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
72 ≤-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
73 ≤-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
74 ≤-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
75
779
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
76 <-ftrans : {x y z : Ordinal } → x <= y → y <= z → x <= z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
77 <-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
78 <-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
79 <-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
80 <-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
81
770
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 769
diff changeset
82 <=to≤ : {x y : Ordinal } → x <= y → * x ≤ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 769
diff changeset
83 <=to≤ (case1 eq) = case1 (cong (*) eq)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 769
diff changeset
84 <=to≤ (case2 lt) = case2 lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 769
diff changeset
85
779
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
86 ≤to<= : {x y : Ordinal } → * x ≤ * y → x <= y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
87 ≤to<= (case1 eq) = case1 ( subst₂ (λ j k → j ≡ k ) &iso &iso (cong (&) eq) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
88 ≤to<= (case2 lt) = case2 lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
89
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
90 <-irr : {a b : HOD} → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
91 <-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (sym a=b) b<a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
92 <-irr {a} {b} (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
93 (IsStrictPartialOrder.trans PO b<a a<b)
490
00c71d1dc316 IsPartialOrder
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 489
diff changeset
94
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
95 ptrans = IsStrictPartialOrder.trans PO
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
96
492
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
97 open _==_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
98 open _⊆_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
99
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
100 --
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
101 -- Closure of ≤-monotonic function f has total order
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
102 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
103
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
104 ≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set (Level.suc n)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
105 ≤-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x ≤ * (f x) ) ∧ odef A (f x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
106
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
107 data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
108 init : {s1 : Ordinal } → odef A s → s ≡ s1 → FClosure A f s s1
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
109 fsuc : (x : Ordinal) ( p : FClosure A f s x ) → FClosure A f s (f x)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
110
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
111 A∋fc : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
112 A∋fc {A} s f mf (init as refl ) = as
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
113 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) )
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
114
714
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 713
diff changeset
115 A∋fcs : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A s
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
116 A∋fcs {A} s f mf (init as refl) = as
714
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 713
diff changeset
117 A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 713
diff changeset
118
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
119 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → * s ≤ * y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
120 s≤fc {A} s {.s} f mf (init x refl ) = case1 refl
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
121 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
122 ... | case1 x=fx = subst (λ k → * s ≤ * k ) (*≡*→≡ x=fx) ( s≤fc {A} s f mf fcy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
123 ... | case2 x<fx with s≤fc {A} s f mf fcy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
124 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym s≡x) refl x<fx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
125 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx )
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
126
557
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
127
559
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
128 fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
129 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → Tri (* x < * y) (* x ≡ * y) (* y < * x )
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
130 fcn-cmp-1 : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
131 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → * x < * y → (* (f x) ≡ * y) ∨ ( * (f x) < * y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
132 fcn-cmp-1 {A} s f mf (init x refl) (init x₁ refl) x<y = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
133 fcn-cmp-1 {A} s f mf (init x refl) (fsuc x₁ cy) x<y = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
134 fcn-cmp-1 {A} s f mf (fsuc x cx) (init ay refl) x<y = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
135 fcn-cmp-1 {A} s f mf (fsuc x cx) (fsuc y cy) x<y with proj1 (mf x (A∋fc s f mf cx)) | proj1 (mf y (A∋fc s f mf cy))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
136 ... | case1 eqx | case1 eqy = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
137 ... | case1 eqx | case2 lt = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
138 ... | case2 lt | case1 eqy = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
139 ... | case2 ltx | case2 lty = {!!}
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
140
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
141 fcn-cmp {A} s {.s} {y} f mf (init ax refl) (init ay refl) = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
142 fcn-cmp {A} s {.s} {.(f x)} f mf (init ax refl) (fsuc x cy) = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
143 fcn-cmp {A} s {.(f x)} {y} f mf (fsuc x cx) (init ay refl) = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
144 fcn-cmp {A} s {.(f x)} {.(f y)} f mf (fsuc x cx) (fsuc y cy) with proj1 (mf x (A∋fc s f mf cx)) | proj1 (mf y (A∋fc s f mf cy))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
145 ... | case1 eqx | case1 eqy = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
146 ... | case1 eqx | case2 lt = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
147 ... | case2 lt | case1 eqy = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
148 ... | case2 ltx | case2 lty = {!!}
563
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
149
729
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 728
diff changeset
150
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
151 -- open import Relation.Binary.Properties.Poset as Poset
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
152
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
153 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
154 IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b) → Tri (a < b) (a ≡ b) (b < a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
155
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
156 ⊆-IsTotalOrderSet : { A B : HOD } → B ⊆ A → IsTotalOrderSet A → IsTotalOrderSet B
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
157 ⊆-IsTotalOrderSet {A} {B} B⊆A T ax ay = T (incl B⊆A ax) (incl B⊆A ay)
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
158
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
159 _⊆'_ : ( A B : HOD ) → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
160 _⊆'_ A B = {x : Ordinal } → odef A x → odef B x
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
161
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
162 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
163 -- inductive maxmum tree from x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
164 -- tree structure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
165 --
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
166
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
167 record HasPrev (A B : HOD) {x : Ordinal } (xa : odef A x) ( f : Ordinal → Ordinal ) : Set n where
533
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
168 field
534
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 533
diff changeset
169 y : Ordinal
541
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 540
diff changeset
170 ay : odef B y
534
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 533
diff changeset
171 x=fy : x ≡ f y
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
172
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
173 record IsSup (A B : HOD) {x : Ordinal } (xa : odef A x) : Set n where
654
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 653
diff changeset
174 field
779
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
175 x<sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x )
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
176
656
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
177 record SUP ( A B : HOD ) : Set (Level.suc n) where
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
178 field
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
179 sup : HOD
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
180 A∋maximal : A ∋ sup
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
181 x<sup : {x : HOD} → B ∋ x → (x ≡ sup ) ∨ (x < sup ) -- B is Total, use positive
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
182
690
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
183 --
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
184 -- sup and its fclosure is in a chain HOD
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
185 -- chain HOD is sorted by sup as Ordinal and <-ordered
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
186 -- whole chain is a union of separated Chain
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
187 -- minimum index is y not ϕ
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
188 --
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
189
787
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 786
diff changeset
190 record ChainP (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal) (u : Ordinal) : Set n where
690
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
191 field
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
192 fcy<sup : {z : Ordinal } → FClosure A f y z → (z ≡ supf u) ∨ ( z << supf u )
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
193 order : {s z1 : Ordinal} → (lt : supf s o< supf u ) → FClosure A f (supf s ) z1 → (z1 ≡ supf u ) ∨ ( z1 << supf u )
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
194 supu=u : supf u ≡ u
694
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
195
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
196 -- Union of supf z which o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
197 --
690
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
198
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
199 data UChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
200 (supf : Ordinal → Ordinal) (x : Ordinal) : (z : Ordinal) → Set n where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
201 ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain A f mf ay supf x z
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
202 ch-is-sup : (u : Ordinal) {z : Ordinal } (u≤x : u o≤ x) ( is-sup : ChainP A f mf ay supf u )
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
203 ( fc : FClosure A f (supf u) z ) → UChain A f mf ay supf x z
694
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
204
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
205 ∈∧P→o< : {A : HOD } {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
206 ∈∧P→o< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
207
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
208 UnionCF : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
209 ( supf : Ordinal → Ordinal ) ( x : Ordinal ) → HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
210 UnionCF A f mf ay supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
211 = record { od = record { def = λ z → odef A z ∧ UChain A f mf ay supf x z } ; odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy }
662
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 661
diff changeset
212
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
213 record ZChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
214 {y : Ordinal} (ay : odef A y) ( z : Ordinal ) : Set (Level.suc n) where
655
b602e3f070df UChain rewrite
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 654
diff changeset
215 field
694
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
216 supf : Ordinal → Ordinal
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
217 chain : HOD
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
218 chain = UnionCF A f mf ay supf z
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
219 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
220 chain⊆A : chain ⊆' A
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
221 chain∋init : odef chain y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
222 initial : {z : Ordinal } → odef chain z → * y ≤ * z
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
223 f-next : {a : Ordinal } → odef chain a → odef chain (f a)
654
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 653
diff changeset
224 f-total : IsTotalOrderSet chain
756
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 755
diff changeset
225
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
226 sup : {x : Ordinal } → x o≤ z → SUP A (UnionCF A f mf ay supf x)
761
9307f895891c edge case done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 760
diff changeset
227 sup=u : {b : Ordinal} → (ab : odef A b) → b o< z → IsSup A (UnionCF A f mf ay supf (osuc b)) ab → supf b ≡ b
795
408e7e8a3797 csupf depends on order cyclicly
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 794
diff changeset
228 supf-is-sup : {x : Ordinal } → (x≤z : x o≤ z) → supf x ≡ & (SUP.sup (sup x≤z) )
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
229 csupf : {b : Ordinal } → b o≤ z → odef (UnionCF A f mf ay supf b) (supf b)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
230 supf≤x :{x : Ordinal } → z o≤ x → supf z ≡ supf x
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
231
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
232 fcy<sup : {u w : Ordinal } → u o≤ z → FClosure A f y w → (w ≡ supf u ) ∨ ( w << supf u ) -- different from order because y o< supf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
233 fcy<sup {u} {w} u≤z fc with SUP.x<sup (sup u≤z) ⟪ subst (λ k → odef A k ) (sym &iso) (A∋fc {A} y f mf fc)
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
234 , ch-init (subst (λ k → FClosure A f y k) (sym &iso) fc ) ⟫
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
235 ... | case1 eq = case1 (subst (λ k → k ≡ supf u ) &iso (trans (cong (&) eq) (sym (supf-is-sup u≤z ) ) ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
236 ... | case2 lt = case2 (subst (λ k → * w < k ) (subst (λ k → k ≡ _ ) *iso (cong (*) (sym (supf-is-sup u≤z ))) ) lt )
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
237
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
238 supf-mono : {x y1 : Ordinal } → x o< y1 → supf x o≤ supf y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
239 supf-mono {x} {y1} x<y1 = sf<sy1 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
240 -- supf x << supf y1 → supf x o< supf y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
241 -- x o< y1 → supf x <= supf y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
242 -- z o≤ x → supf x ≡ supf y1 ≡ supf z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
243 -- x o< z → z o< y1 → supf x ≡ supf y1 ≡ supf z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
244 supy : {x : Ordinal } → x o≤ z → FClosure A f y (supf x) → supf x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
245 supy {x} x≤z fcx = ? where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
246 zc06 : ( * y ≡ SUP.sup (sup x≤z )) ∨ ( * y < SUP.sup ( sup x≤z ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
247 zc06 = SUP.x<sup (sup x≤z) ⟪ subst (λ k → odef A k ) (sym &iso) ay , ch-init (init ay (sym &iso) ) ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
248 sf<sy1 : supf x o≤ supf y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
249 sf<sy1 with trio< x z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
250 ... | tri> ¬a ¬b c = o≤-refl0 (( trans (sym (supf≤x (o<→≤ c))) (supf≤x (ordtrans (ordtrans c x<y1 ) <-osuc ) ) ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
251 ... | tri≈ ¬a b ¬c = o≤-refl0 (trans (sym (supf≤x (o≤-refl0 (sym b)))) (supf≤x (subst (λ k → k o< osuc y1) b (o<→≤ x<y1))))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
252 ... | tri< x<z ¬b ¬c with trio< (supf x) (supf y1)
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
253 ... | tri< a ¬b ¬c = o<→≤ a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
254 ... | tri≈ ¬a b ¬c = o≤-refl0 b
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
255 ... | tri> ¬a ¬b sy1<sx with trio< z y1
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
256 ... | tri< a ¬b ¬c = ?
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
257 ... | tri≈ ¬a b ¬c = ?
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
258 ... | tri> ¬a ¬b y1<z = ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
259 zc04 : x o< z → y1 o≤ z → supf x o≤ supf y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
260 zc04 x<z y1≤z with csupf (o<→≤ x<z) | csupf y1≤z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
261 ... | ⟪ ax , ch-init fcx ⟫ | ⟪ ay1 , ch-init fcy1 ⟫ with fcy<sup (o<→≤ x<z) fcy1
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
262 ... | case1 eq = o≤-refl0 (sym eq)
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
263 ... | case2 lt with fcy<sup y1≤z fcx
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
264 ... | case1 eq = o≤-refl0 eq
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
265 ... | case2 lt1 = ⊥-elim ( <-irr (case2 lt) lt1 )
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
266 zc04 x<z y1≤z | ⟪ ax , ch-is-sup ux ux≤z is-sup-x fcx ⟫ | ⟪ ay1 , ch-init fcy1 ⟫ with fcy<sup (o<→≤ x<z) fcy1
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
267 ... | case1 eq = o≤-refl0 (sym eq)
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
268 ... | case2 lt with trio< (supf x) (supf y1)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
269 ... | tri< a ¬b ¬c = o<→≤ a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
270 ... | tri≈ ¬a b ¬c = o≤-refl0 b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
271 ... | tri> ¬a ¬b y1<z = ? where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
272 zc05 : ( supf y1 ≡ supf ux ) ∨ (supf y1 << supf ux )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
273 zc05 = ChainP.fcy<sup is-sup-x fcy1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
274 zc04 x<z y1≤z | ⟪ ax , ch-is-sup ux ux≤z is-sup-x fcx ⟫ | ⟪ ay1 , ch-init fcy1 ⟫ | case2 lt1 = ? -- sy1 << sx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
275 zc04 x<z y1≤z | ⟪ ax , ch-init fcx ⟫ | ⟪ ay1 , ch-is-sup uy1 uy1≤z is-sup-y1 fcy1 ⟫ = ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
276 zc04 x<z y1≤z | ⟪ ax , ch-is-sup ux ux≤z is-sup-x fcx ⟫ | ⟪ ay1 , ch-is-sup uy1 uy1≤z is-sup-y1 fcy1 ⟫ = ?
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
277 -- ... | tri< a ¬b ¬c = csupf (o<→≤ a)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
278 -- ... | tri≈ ¬a b ¬c = csupf (o≤-refl0 b)
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
279 -- ... | tri> ¬a ¬b c = subst (λ k → odef (UnionCF A f mf ay1 supf x) k ) ? (csupf ? )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
280 -- csy1 : odef (UnionCF A f mf ay1 supf y1) (supf y1)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
281 -- csy1 = csupf ?
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
282
784
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 783
diff changeset
283 order : {b s z1 : Ordinal} → b o< z → supf s o< supf b → FClosure A f (supf s) z1 → (z1 ≡ supf b) ∨ (z1 << supf b)
785
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
284 order {b} {s} {z1} b<z sf<sb fc = zc04 where
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
285 zc01 : {z1 : Ordinal } → FClosure A f (supf s) z1 → UnionCF A f mf ay supf b ∋ * z1
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
286 zc01 (init x refl ) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc03 where
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
287 s<b : s o< b
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
288 s<b with trio< s b
785
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
289 ... | tri< a ¬b ¬c = a
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
290 ... | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (cong supf b) sf<sb )
785
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
291 ... | tri> ¬a ¬b c with osuc-≡< ( supf-mono c )
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
292 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sf<sb )
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
293 ... | case2 lt = ⊥-elim ( o<> lt sf<sb )
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
294 s<z : s o< z
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
295 s<z = ordtrans s<b b<z
785
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
296 zc03 : odef (UnionCF A f mf ay supf b) (supf s)
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
297 zc03 with csupf (o<→≤ s<z )
785
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
298 ... | ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
299 ... | ⟪ as , ch-is-sup u u≤x is-sup fc ⟫ = ⟪ as , ch-is-sup u (ordtrans u≤x (osucc s<b)) is-sup fc ⟫
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
300 zc01 (fsuc x fc) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc04 where
785
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
301 zc04 : odef (UnionCF A f mf ay supf b) (f x)
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
302 zc04 with subst (λ k → odef (UnionCF A f mf ay supf b) k ) &iso (zc01 fc )
7472e3dc002b order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 784
diff changeset
303 ... | ⟪ as , ch-init fc ⟫ = ⟪ proj2 (mf _ as) , ch-init (fsuc _ fc) ⟫
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
304 ... | ⟪ as , ch-is-sup u u≤x is-sup fc ⟫ = ⟪ proj2 (mf _ as) , ch-is-sup u u≤x is-sup (fsuc _ fc) ⟫
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
305 zc00 : ( * z1 ≡ SUP.sup (sup (o<→≤ b<z) )) ∨ ( * z1 < SUP.sup ( sup (o<→≤ b<z) ) )
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
306 zc00 = SUP.x<sup (sup (o<→≤ b<z)) (zc01 fc )
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
307 zc04 : (z1 ≡ supf b) ∨ (z1 << supf b)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
308 zc04 with zc00
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
309 ... | case1 eq = case1 (subst₂ (λ j k → j ≡ k ) &iso (sym (supf-is-sup (o<→≤ b<z) ) ) (cong (&) eq) )
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
310 ... | case2 lt = case2 (subst₂ (λ j k → j < k ) refl (subst₂ (λ j k → j ≡ k ) *iso refl (cong (*) (sym (supf-is-sup (o<→≤ b<z) ) ))) lt )
756
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 755
diff changeset
311
653
4186c0331abb sind again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 638
diff changeset
312
728
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
313 record ZChain1 ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
314 {y : Ordinal} (ay : odef A y) (zc : ZChain A f mf ay (& A)) ( z : Ordinal ) : Set (Level.suc n) where
728
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
315 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
316 is-max : {a b : Ordinal } → (ca : odef (UnionCF A f mf ay (ZChain.supf zc) z) a ) → b o< z → (ab : odef A b)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
317 → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) z) ab f ∨ IsSup A (UnionCF A f mf ay (ZChain.supf zc) z) ab
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
318 → * a < * b → odef ((UnionCF A f mf ay (ZChain.supf zc) z)) b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
319
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
320 record Maximal ( A : HOD ) : Set (Level.suc n) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
321 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
322 maximal : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
323 A∋maximal : A ∋ maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
324 ¬maximal<x : {x : HOD} → A ∋ x → ¬ maximal < x -- A is Partial, use negative
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
325
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
326 -- data UChain is total
684
822fce8af579 no transfinite on data Chain trichotomos
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 683
diff changeset
327
694
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
328 chain-total : (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
329 {s s1 a b : Ordinal } ( ca : UChain A f mf ay supf s a ) ( cb : UChain A f mf ay supf s1 b ) → Tri (* a < * b) (* a ≡ * b) (* b < * a )
694
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
330 chain-total A f mf {y} ay supf {xa} {xb} {a} {b} ca cb = ct-ind xa xb ca cb where
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
331 ct-ind : (xa xb : Ordinal) → {a b : Ordinal} → UChain A f mf ay supf xa a → UChain A f mf ay supf xb b → Tri (* a < * b) (* a ≡ * b) (* b < * a)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
332 ct-ind xa xb {a} {b} (ch-init fca) (ch-init fcb) = fcn-cmp y f mf fca fcb
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
333 ct-ind xa xb {a} {b} (ch-init fca) (ch-is-sup ub u≤x supb fcb) with ChainP.fcy<sup supb fca
766
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
334 ... | case1 eq with s≤fc (supf ub) f mf fcb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
335 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
336 ct00 : * a ≡ * b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
337 ct00 = trans (cong (*) eq) eq1
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
338 ... | case2 lt = tri< ct01 (λ eq → <-irr (case1 (sym eq)) ct01) (λ lt → <-irr (case2 ct01) lt) where
766
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
339 ct01 : * a < * b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
340 ct01 = subst (λ k → * k < * b ) (sym eq) lt
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
341 ct-ind xa xb {a} {b} (ch-init fca) (ch-is-sup ub u≤x supb fcb) | case2 lt = tri< ct01 (λ eq → <-irr (case1 (sym eq)) ct01) (λ lt → <-irr (case2 ct01) lt) where
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
342 ct00 : * a < * (supf ub)
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
343 ct00 = lt
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
344 ct01 : * a < * b
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
345 ct01 with s≤fc (supf ub) f mf fcb
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
346 ... | case1 eq = subst (λ k → * a < k ) eq ct00
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
347 ... | case2 lt = IsStrictPartialOrder.trans POO ct00 lt
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
348 ct-ind xa xb {a} {b} (ch-is-sup ua u≤x supa fca) (ch-init fcb) with ChainP.fcy<sup supa fcb
766
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
349 ... | case1 eq with s≤fc (supf ua) f mf fca
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
350 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
351 ct00 : * a ≡ * b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
352 ct00 = sym (trans (cong (*) eq) eq1 )
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
353 ... | case2 lt = tri> (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01 where
766
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
354 ct01 : * b < * a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
355 ct01 = subst (λ k → * k < * a ) (sym eq) lt
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
356 ct-ind xa xb {a} {b} (ch-is-sup ua u≤x supa fca) (ch-init fcb) | case2 lt = tri> (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01 where
749
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
357 ct00 : * b < * (supf ua)
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
358 ct00 = lt
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
359 ct01 : * b < * a
749
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
360 ct01 with s≤fc (supf ua) f mf fca
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
361 ... | case1 eq = subst (λ k → * b < k ) eq ct00
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
362 ... | case2 lt = IsStrictPartialOrder.trans POO ct00 lt
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
363 ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) with trio< (supf ua) (supf ub)
775
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 774
diff changeset
364 ... | tri< a₁ ¬b ¬c with ChainP.order supb a₁ fca
766
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
365 ... | case1 eq with s≤fc (supf ub) f mf fcb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
366 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
367 ct00 : * a ≡ * b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
368 ct00 = trans (cong (*) eq) eq1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
369 ... | case2 lt = tri< ct02 (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
370 ct02 : * a < * b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
371 ct02 = subst (λ k → * k < * b ) (sym eq) lt
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
372 ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri< a₁ ¬b ¬c | case2 lt = tri< ct02 (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt) where
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
373 ct03 : * a < * (supf ub)
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
374 ct03 = lt
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
375 ct02 : * a < * b
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
376 ct02 with s≤fc (supf ub) f mf fcb
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
377 ... | case1 eq = subst (λ k → * a < k ) eq ct03
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
378 ... | case2 lt = IsStrictPartialOrder.trans POO ct03 lt
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
379 ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri≈ ¬a eq ¬c
769
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 768
diff changeset
380 = fcn-cmp (supf ua) f mf fca (subst (λ k → FClosure A f k b ) (sym eq) fcb )
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
381 ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri> ¬a ¬b c with ChainP.order supa c fcb
766
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
382 ... | case1 eq with s≤fc (supf ua) f mf fca
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
383 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
384 ct00 : * a ≡ * b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
385 ct00 = sym (trans (cong (*) eq) eq1)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
386 ... | case2 lt = tri> (λ lt → <-irr (case2 ct02) lt) (λ eq → <-irr (case1 eq) ct02) ct02 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
387 ct02 : * b < * a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 765
diff changeset
388 ct02 = subst (λ k → * k < * a ) (sym eq) lt
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
389 ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri> ¬a ¬b c | case2 lt = tri> (λ lt → <-irr (case2 ct04) lt) (λ eq → <-irr (case1 (eq)) ct04) ct04 where
749
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
390 ct05 : * b < * (supf ua)
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
391 ct05 = lt
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
392 ct04 : * b < * a
749
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
393 ct04 with s≤fc (supf ua) f mf fca
689
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
394 ... | case1 eq = subst (λ k → * b < k ) eq ct05
34650e39e553 Chain is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 688
diff changeset
395 ... | case2 lt = IsStrictPartialOrder.trans POO ct05 lt
684
822fce8af579 no transfinite on data Chain trichotomos
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 683
diff changeset
396
743
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 742
diff changeset
397 init-uchain : (A : HOD) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } → (ay : odef A y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 742
diff changeset
398 { supf : Ordinal → Ordinal } { x : Ordinal } → odef (UnionCF A f mf ay supf x) y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
399 init-uchain A f mf ay = ⟪ ay , ch-init (init ay refl) ⟫
743
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 742
diff changeset
400
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
401 Zorn-lemma : { A : HOD }
464
5acf6483a9e3 Zorn lemma start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 447
diff changeset
402 → o∅ o< & A
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
403 → ( ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
404 → Maximal A
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
405 Zorn-lemma {A} 0<A supP = zorn00 where
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
406 <-irr0 : {a b : HOD} → A ∋ a → A ∋ b → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
407 <-irr0 {a} {b} A∋a A∋b = <-irr
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
408 z07 : {y : Ordinal} {A : HOD } → {P : Set n} → odef A y ∧ P → y o< & A
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
409 z07 {y} {A} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
760
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 759
diff changeset
410 z09 : {b : Ordinal } { A : HOD } → odef A b → b o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 759
diff changeset
411 z09 {b} {A} ab = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
412 s : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
413 s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A ))
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
414 as : A ∋ * ( & s )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
415 as = subst (λ k → odef A (& k) ) (sym *iso) ( ODC.x∋minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) )
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
416 as0 : odef A (& s )
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
417 as0 = subst (λ k → odef A k ) &iso as
547
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 546
diff changeset
418 s<A : & s o< & A
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
419 s<A = c<→o< (subst (λ k → odef A (& k) ) *iso as )
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
420 HasMaximal : HOD
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
421 HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) → odef A m → ¬ (* x < * m)) } ; odmax = & A ; <odmax = z07 }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
422 no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → odef A x ∧ ((m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m) )) → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
423 no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
424 Gtx : { x : HOD} → A ∋ x → HOD
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
425 Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
426 z08 : ¬ Maximal A → HasMaximal =h= od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
427 z08 nmx = record { eq→ = λ {x} lt → ⊥-elim ( nmx record {maximal = * x ; A∋maximal = subst (λ k → odef A k) (sym &iso) (proj1 lt)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
428 ; ¬maximal<x = λ {y} ay → subst (λ k → ¬ (* x < k)) *iso (proj2 lt (& y) ay) } ) ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
429 x-is-maximal : ¬ Maximal A → {x : Ordinal} → (ax : odef A x) → & (Gtx (subst (λ k → odef A k ) (sym &iso) ax)) ≡ o∅ → (m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
430 x-is-maximal nmx {x} ax nogt m am = ⟪ subst (λ k → odef A k) &iso (subst (λ k → odef A k ) (sym &iso) ax) , ¬x<m ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
431 ¬x<m : ¬ (* x < * m)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
432 ¬x<m x<m = ∅< {Gtx (subst (λ k → odef A k ) (sym &iso) ax)} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫ (≡o∅→=od∅ nogt)
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
433
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
434 -- Uncountable ascending chain by axiom of choice
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
435 cf : ¬ Maximal A → Ordinal → Ordinal
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
436 cf nmx x with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
437 ... | no _ = o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
438 ... | yes ax with is-o∅ (& ( Gtx ax ))
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
439 ... | yes nogt = -- no larger element, so it is maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
440 ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
441 ... | no not = & (ODC.minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq)))
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
442 is-cf : (nmx : ¬ Maximal A ) → {x : Ordinal} → odef A x → odef A (cf nmx x) ∧ ( * x < * (cf nmx x) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
443 is-cf nmx {x} ax with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
444 ... | no not = ⊥-elim ( not (subst (λ k → odef A k ) (sym &iso) ax ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
445 ... | yes ax with is-o∅ (& ( Gtx ax ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
446 ... | yes nogt = ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
447 ... | no not = ODC.x∋minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq))
606
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
448
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
449 ---
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
450 --- infintie ascention sequence of f
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
451 ---
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
452 cf-is-<-monotonic : (nmx : ¬ Maximal A ) → (x : Ordinal) → odef A x → ( * x < * (cf nmx x) ) ∧ odef A (cf nmx x )
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
453 cf-is-<-monotonic nmx x ax = ⟪ proj2 (is-cf nmx ax ) , proj1 (is-cf nmx ax ) ⟫
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
454 cf-is-≤-monotonic : (nmx : ¬ Maximal A ) → ≤-monotonic-f A ( cf nmx )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
455 cf-is-≤-monotonic nmx x ax = ⟪ case2 (proj1 ( cf-is-<-monotonic nmx x ax )) , proj2 ( cf-is-<-monotonic nmx x ax ) ⟫
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
456
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
457 chain-mono : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
458 {a b c : Ordinal} → a o≤ b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
459 → odef (UnionCF A f mf ay supf a) c → odef (UnionCF A f mf ay supf b) c
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
460 chain-mono f mf ay supf {a} {b} {c} a≤b ⟪ ua , ch-init fc ⟫ =
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
461 ⟪ ua , ch-init fc ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
462 chain-mono f mf ay supf {a} {b} {c} a≤b ⟪ uaa , ch-is-sup ua ua<x is-sup fc ⟫ =
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
463 ⟪ uaa , ch-is-sup ua (ordtrans<-≤ ua<x (osucc a≤b )) is-sup fc ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
464
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
465 sp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (zc : ZChain A f mf as0 (& A) )
653
4186c0331abb sind again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 638
diff changeset
466 (total : IsTotalOrderSet (ZChain.chain zc) ) → SUP A (ZChain.chain zc)
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
467 sp0 f mf zc total = supP (ZChain.chain zc) (ZChain.chain⊆A zc) total
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
468 zc< : {x y z : Ordinal} → {P : Set n} → (x o< y → P) → x o< z → z o< y → P
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
469 zc< {x} {y} {z} {P} prev x<z z<y = prev (ordtrans x<z z<y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
470
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
471 SZ1 : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)
728
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
472 {init : Ordinal} (ay : odef A init) (zc : ZChain A f mf ay (& A)) (x : Ordinal) → ZChain1 A f mf ay zc x
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
473 SZ1 f mf {y} ay zc x = TransFinite { λ x → ZChain1 A f mf ay zc x } zc1 x where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
474 chain-mono1 : {a b c : Ordinal} → a o≤ b
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
475 → odef (UnionCF A f mf ay (ZChain.supf zc) a) c → odef (UnionCF A f mf ay (ZChain.supf zc) b) c
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
476 chain-mono1 {a} {b} {c} a≤b = chain-mono f mf ay (ZChain.supf zc) a≤b
735
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 734
diff changeset
477 is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 734
diff changeset
478 b o< x → (ab : odef A b) →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 734
diff changeset
479 HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) ab f →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 734
diff changeset
480 * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b
749
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
481 is-max-hp x {a} {b} ua b<x ab has-prev a<b with HasPrev.ay has-prev
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
482 ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
483 ... | ⟪ ab0 , ch-is-sup u u≤x is-sup fc ⟫ = ⟪ ab ,
749
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
484 subst (λ k → UChain A f mf ay (ZChain.supf zc) x k )
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
485 (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u≤x is-sup (fsuc _ fc)) ⟫
728
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
486 zc1 : (x : Ordinal) → ((y₁ : Ordinal) → y₁ o< x → ZChain1 A f mf ay zc y₁) → ZChain1 A f mf ay zc x
732
ddeb107b6f71 bchain can be reached from upwords by f. so it is worng.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 729
diff changeset
487 zc1 x prev with Oprev-p x
756
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 755
diff changeset
488 ... | yes op = record { is-max = is-max } where
732
ddeb107b6f71 bchain can be reached from upwords by f. so it is worng.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 729
diff changeset
489 px = Oprev.oprev op
789
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 788
diff changeset
490 zc-b<x : {b : Ordinal } → b o< x → b o< osuc px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 788
diff changeset
491 zc-b<x {b} lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt
728
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
492 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
493 b o< x → (ab : odef A b) →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
494 HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) ab f ∨ IsSup A (UnionCF A f mf ay (ZChain.supf zc) x) ab →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
495 * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b
735
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 734
diff changeset
496 is-max {a} {b} ua b<x ab (case1 has-prev) a<b = is-max-hp x {a} {b} ua b<x ab has-prev a<b
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
497 is-max {a} {b} ua b<x ab (case2 is-sup) a<b
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
498 = ⟪ ab , ch-is-sup b (o<→≤ b<x) m06 (subst (λ k → FClosure A f k b) m05 (init ab refl)) ⟫ where
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
499 b<A : b o< & A
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
500 b<A = z09 ab
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
501 m05 : b ≡ ZChain.supf zc b
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
502 m05 = sym ( ZChain.sup=u zc ab (z09 ab)
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
503 record { x<sup = λ {z} uz → IsSup.x<sup is-sup (chain-mono1 (osucc b<x) uz ) } )
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
504 m08 : {z : Ordinal} → (fcz : FClosure A f y z ) → z <= ZChain.supf zc b
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
505 m08 {z} fcz = ZChain.fcy<sup zc (o<→≤ b<A) fcz
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
506 m09 : {sup1 z1 : Ordinal} → (ZChain.supf zc sup1) o< (ZChain.supf zc b)
769
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 768
diff changeset
507 → FClosure A f (ZChain.supf zc sup1) z1 → z1 <= ZChain.supf zc b
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
508 m09 {sup1} {z} s<b fcz = ZChain.order zc b<A s<b fcz
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
509 m06 : ChainP A f mf ay (ZChain.supf zc) b
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
510 m06 = record { fcy<sup = m08 ; order = m09 ; supu=u = ZChain.sup=u zc ab b<A {!!} }
756
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 755
diff changeset
511 ... | no lim = record { is-max = is-max } where
734
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 733
diff changeset
512 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 733
diff changeset
513 b o< x → (ab : odef A b) →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 733
diff changeset
514 HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) ab f ∨ IsSup A (UnionCF A f mf ay (ZChain.supf zc) x) ab →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 733
diff changeset
515 * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b
735
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 734
diff changeset
516 is-max {a} {b} ua b<x ab (case1 has-prev) a<b = is-max-hp x {a} {b} ua b<x ab has-prev a<b
743
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 742
diff changeset
517 is-max {a} {b} ua b<x ab (case2 is-sup) a<b with IsSup.x<sup is-sup (init-uchain A f mf ay )
789
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 788
diff changeset
518 ... | case1 b=y = ⟪ subst (λ k → odef A k ) b=y ay , ch-init (subst (λ k → FClosure A f y k ) b=y (init ay refl )) ⟫
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
519 ... | case2 y<b = chain-mono1 (osucc b<x)
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
520 ⟪ ab , ch-is-sup b (ordtrans o≤-refl <-osuc ) m06 (subst (λ k → FClosure A f k b) m05 (init ab refl)) ⟫ where
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
521 m09 : b o< & A
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
522 m09 = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
523 m07 : {z : Ordinal} → FClosure A f y z → z <= ZChain.supf zc b
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
524 m07 {z} fc = ZChain.fcy<sup zc (o<→≤ m09) fc
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
525 m08 : {sup1 z1 : Ordinal} → (ZChain.supf zc sup1) o< (ZChain.supf zc b)
769
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 768
diff changeset
526 → FClosure A f (ZChain.supf zc sup1) z1 → z1 <= ZChain.supf zc b
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
527 m08 {sup1} {z1} s<b fc = ZChain.order zc m09 s<b fc
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
528 m05 : b ≡ ZChain.supf zc b
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
529 m05 = sym (ZChain.sup=u zc ab m09
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
530 record { x<sup = λ lt → IsSup.x<sup is-sup (chain-mono1 (osucc b<x) lt )} ) -- ZChain on x
790
201b66da4e69 remove unnesesary part in SZ1 the second TransFinite induction for is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 789
diff changeset
531 m06 : ChainP A f mf ay (ZChain.supf zc) b
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
532 m06 = record { fcy<sup = m07 ; order = m08 ; supu=u = ZChain.sup=u zc ab m09 {!!} }
727
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 726
diff changeset
533
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
534 ---
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
535 --- the maximum chain has fix point of any ≤-monotonic function
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
536 ---
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
537 fixpoint : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (zc : ZChain A f mf as0 (& A) )
633
6cd4a483122c ZChain1 is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 631
diff changeset
538 → (total : IsTotalOrderSet (ZChain.chain zc) )
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
539 → f (& (SUP.sup (sp0 f mf zc total ))) ≡ & (SUP.sup (sp0 f mf zc total))
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
540 fixpoint f mf zc total = z14 where
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
541 chain = ZChain.chain zc
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
542 sp1 = sp0 f mf zc total
712
92275389e623 fix is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 711
diff changeset
543 z10 : {a b : Ordinal } → (ca : odef chain a ) → b o< & A → (ab : odef A b )
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
544 → HasPrev A chain ab f ∨ IsSup A chain {b} ab -- (supO chain (ZChain.chain⊆A zc) (ZChain.f-total zc) ≡ b )
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
545 → * a < * b → odef chain b
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
546 z10 = ZChain1.is-max (SZ1 f mf as0 zc (& A) )
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
547 z11 : & (SUP.sup sp1) o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
548 z11 = c<→o< ( SUP.A∋maximal sp1)
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
549 z12 : odef chain (& (SUP.sup sp1))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
550 z12 with o≡? (& s) (& (SUP.sup sp1))
653
4186c0331abb sind again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 638
diff changeset
551 ... | yes eq = subst (λ k → odef chain k) eq ( ZChain.chain∋init zc )
712
92275389e623 fix is-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 711
diff changeset
552 ... | no ne = z10 {& s} {& (SUP.sup sp1)} ( ZChain.chain∋init zc ) z11 (SUP.A∋maximal sp1)
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
553 (case2 z19 ) z13 where
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
554 z13 : * (& s) < * (& (SUP.sup sp1))
653
4186c0331abb sind again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 638
diff changeset
555 z13 with SUP.x<sup sp1 ( ZChain.chain∋init zc )
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
556 ... | case1 eq = ⊥-elim ( ne (cong (&) eq) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
557 ... | case2 lt = subst₂ (λ j k → j < k ) (sym *iso) (sym *iso) lt
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
558 z19 : IsSup A chain {& (SUP.sup sp1)} (SUP.A∋maximal sp1)
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
559 z19 = record { x<sup = z20 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
560 z20 : {y : Ordinal} → odef chain y → (y ≡ & (SUP.sup sp1)) ∨ (y << & (SUP.sup sp1))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
561 z20 {y} zy with SUP.x<sup sp1 (subst (λ k → odef chain k ) (sym &iso) zy)
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
562 ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso ( cong (&) y=p ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
563 ... | case2 y<p = case2 (subst (λ k → * y < k ) (sym *iso) y<p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
564 -- λ {y} zy → subst (λ k → (y ≡ & k ) ∨ (y << & k)) ? (SUP.x<sup sp1 ? ) }
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
565 z14 : f (& (SUP.sup (sp0 f mf zc total ))) ≡ & (SUP.sup (sp0 f mf zc total ))
633
6cd4a483122c ZChain1 is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 631
diff changeset
566 z14 with total (subst (λ k → odef chain k) (sym &iso) (ZChain.f-next zc z12 )) z12
631
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
567 ... | tri< a ¬b ¬c = ⊥-elim z16 where
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
568 z16 : ⊥
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
569 z16 with proj1 (mf (& ( SUP.sup sp1)) ( SUP.A∋maximal sp1 ))
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
570 ... | case1 eq = ⊥-elim (¬b (subst₂ (λ j k → j ≡ k ) refl *iso (sym eq) ))
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
571 ... | case2 lt = ⊥-elim (¬c (subst₂ (λ j k → k < j ) refl *iso lt ))
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
572 ... | tri≈ ¬a b ¬c = subst ( λ k → k ≡ & (SUP.sup sp1) ) &iso ( cong (&) b )
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
573 ... | tri> ¬a ¬b c = ⊥-elim z17 where
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
574 z15 : (* (f ( & ( SUP.sup sp1 ))) ≡ SUP.sup sp1) ∨ (* (f ( & ( SUP.sup sp1 ))) < SUP.sup sp1)
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
575 z15 = SUP.x<sup sp1 (subst (λ k → odef chain k ) (sym &iso) (ZChain.f-next zc z12 ))
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
576 z17 : ⊥
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
577 z17 with z15
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
578 ... | case1 eq = ¬b eq
1150b006059b ... give up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 630
diff changeset
579 ... | case2 lt = ¬a lt
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
580
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
581 -- ZChain contradicts ¬ Maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
582 --
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
583 -- ZChain forces fix point on any ≤-monotonic function (fixpoint)
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
584 -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
585 --
697
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 696
diff changeset
586 z04 : (nmx : ¬ Maximal A )
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
587 → (zc : ZChain A (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A))
664
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 663
diff changeset
588 → IsTotalOrderSet (ZChain.chain zc) → ⊥
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
589 z04 nmx zc total = <-irr0 {* (cf nmx c)} {* c} (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (SUP.A∋maximal sp1 ))))
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
590 (subst (λ k → odef A (& k)) (sym *iso) (SUP.A∋maximal sp1) )
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
591 (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) zc total ))) -- x ≡ f x ̄
633
6cd4a483122c ZChain1 is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 631
diff changeset
592 (proj1 (cf-is-<-monotonic nmx c (SUP.A∋maximal sp1 ))) where -- x < f x
6cd4a483122c ZChain1 is not strictly positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 631
diff changeset
593 sp1 : SUP A (ZChain.chain zc)
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
594 sp1 = sp0 (cf nmx) (cf-is-≤-monotonic nmx) zc total
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
595 c = & (SUP.sup sp1)
548
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 547
diff changeset
596
757
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
597 uchain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
598 uchain f mf {y} ay = record { od = record { def = λ x → FClosure A f y x } ; odmax = & A ; <odmax =
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
599 λ {z} cz → subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (A∋fc y f mf cz ))) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
600
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
601 utotal : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
602 → IsTotalOrderSet (uchain f mf ay)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
603 utotal f mf {y} ay {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
604 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
605 uz01 = fcn-cmp y f mf ca cb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
606
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
607 ysup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
608 → SUP A (uchain f mf ay)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
609 ysup f mf {y} ay = supP (uchain f mf ay) (λ lt → A∋fc y f mf lt) (utotal f mf ay)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
610
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
611 initChain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → ZChain A f mf ay o∅
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
612 initChain f mf {y} ay = record { supf = isupf ; chain⊆A = λ lt → proj1 lt ; chain∋init = cy ; sup = {!!} ; supf-is-sup = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
613 ; initial = isy ; f-next = inext ; f-total = itotal ; sup=u = λ _ b<0 → ⊥-elim (¬x<0 b<0) ; supf-mono = mono ; csupf = {!!} } where
764
bbf12d61143f < is wrong
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 763
diff changeset
614 spi = & (SUP.sup (ysup f mf ay))
711
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
615 isupf : Ordinal → Ordinal
768
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 767
diff changeset
616 isupf z = spi
763
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 762
diff changeset
617 sp = ysup f mf ay
767
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
618 asi = SUP.A∋maximal sp
711
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
619 cy : odef (UnionCF A f mf ay isupf o∅) y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
620 cy = ⟪ ay , ch-init (init ay refl) ⟫
759
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 758
diff changeset
621 y<sup : * y ≤ SUP.sup (ysup f mf ay)
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
622 y<sup = SUP.x<sup (ysup f mf ay) (subst (λ k → FClosure A f y k ) (sym &iso) (init ay refl))
786
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 785
diff changeset
623 sup : {x : Ordinal} → x o< o∅ → SUP A (UnionCF A f mf ay isupf x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 785
diff changeset
624 sup {x} lt = ⊥-elim ( ¬x<0 lt )
711
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
625 isy : {z : Ordinal } → odef (UnionCF A f mf ay isupf o∅) z → * y ≤ * z
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
626 isy {z} ⟪ az , uz ⟫ with uz
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
627 ... | ch-init fc = s≤fc y f mf fc
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
628 ... | ch-is-sup u u≤x is-sup fc = ≤-ftrans (subst (λ k → * y ≤ k) (sym *iso) y<sup) (s≤fc (& (SUP.sup (ysup f mf ay))) f mf fc )
711
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
629 inext : {a : Ordinal} → odef (UnionCF A f mf ay isupf o∅) a → odef (UnionCF A f mf ay isupf o∅) (f a)
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
630 inext {a} ua with (proj2 ua)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
631 ... | ch-init fc = ⟪ proj2 (mf _ (proj1 ua)) , ch-init (fsuc _ fc ) ⟫
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
632 ... | ch-is-sup u u≤x is-sup fc = ⟪ proj2 (mf _ (proj1 ua)) , ch-is-sup u u≤x is-sup (fsuc _ fc) ⟫
711
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
633 itotal : IsTotalOrderSet (UnionCF A f mf ay isupf o∅)
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
634 itotal {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
635 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
763
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 762
diff changeset
636 uz01 = chain-total A f mf ay isupf (proj2 ca) (proj2 cb)
786
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 785
diff changeset
637 mono : {x : Ordinal} {z : Ordinal} → x o< z → isupf x o≤ isupf z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 785
diff changeset
638 mono {x} {z} x<z = o≤-refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 785
diff changeset
639 csupf : {z : Ordinal} → z o≤ o∅ → odef (UnionCF A f mf ay isupf z ) (isupf z)
789
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 788
diff changeset
640 csupf {z} z≤0 = ⟪ asi , ch-is-sup o∅ o∅≤z uz02 (init asi refl) ⟫ where
768
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 767
diff changeset
641 uz03 : {z : Ordinal } → FClosure A f y z → (z ≡ isupf spi) ∨ (z << isupf spi)
767
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
642 uz03 {z} fc with SUP.x<sup sp (subst (λ k → FClosure A f y k ) (sym &iso) fc )
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
643 ... | case1 eq = case1 ( begin
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
644 z ≡⟨ sym &iso ⟩
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
645 & (* z) ≡⟨ cong (&) eq ⟩
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
646 spi ∎ ) where open ≡-Reasoning
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
647 ... | case2 lt = case2 (subst (λ k → * z < k ) (sym *iso) lt )
769
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 768
diff changeset
648 uz04 : {sup1 z1 : Ordinal} → isupf sup1 o< isupf spi → FClosure A f (isupf sup1) z1 → (z1 ≡ isupf spi) ∨ (z1 << isupf spi)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 768
diff changeset
649 uz04 {s} {z} s<spi fcz = ⊥-elim ( o<¬≡ refl s<spi )
789
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 788
diff changeset
650 uz02 : ChainP A f mf ay isupf o∅
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
651 uz02 = record { fcy<sup = uz03 ; order = λ {s} {z} → uz04 {s} {z} ; supu=u = ? }
767
6c87cb98abf2 spi <= u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 766
diff changeset
652
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
653 SUP⊆ : { B C : HOD } → B ⊆' C → SUP A C → SUP A B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
654 SUP⊆ {B} {C} B⊆C sup = record { sup = SUP.sup sup ; A∋maximal = SUP.A∋maximal sup ; x<sup = λ lt → SUP.x<sup sup (B⊆C lt) }
711
b1d468186e68 initial chain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 710
diff changeset
655
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
656 --
547
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 546
diff changeset
657 -- create all ZChains under o< x
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
658 --
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
659
674
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 673
diff changeset
660 ind : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal)
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
661 → ((z : Ordinal) → z o< x → ZChain A f mf ay z) → ZChain A f mf ay x
707
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 706
diff changeset
662 ind f mf {y} ay x prev with Oprev-p x
697
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 696
diff changeset
663 ... | yes op = zc4 where
682
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
664 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
665 -- we have previous ordinal to use induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
666 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
667 px = Oprev.oprev op
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
668 zc : ZChain A f mf ay (Oprev.oprev op)
682
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
669 zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
670 px<x : px o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
671 px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc
709
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 708
diff changeset
672 zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 708
diff changeset
673 zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt
697
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 696
diff changeset
674
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
675 pchain : HOD
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
676 pchain = UnionCF A f mf ay (ZChain.supf zc) x
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
677 ptotal : IsTotalOrderSet pchain
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
678 ptotal {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
679 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
680 uz01 = chain-total A f mf ay (ZChain.supf zc) ( (proj2 ca)) ( (proj2 cb))
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
681 pchain⊆A : {y : Ordinal} → odef pchain y → odef A y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
682 pchain⊆A {y} ny = proj1 ny
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
683 pnext : {a : Ordinal} → odef pchain a → odef pchain (f a)
749
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 748
diff changeset
684 pnext {a} ⟪ aa , ch-init fc ⟫ = ⟪ proj2 (mf a aa) , ch-init (fsuc _ fc) ⟫
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
685 pnext {a} ⟪ aa , ch-is-sup u u≤x is-sup fc ⟫ = ⟪ proj2 (mf a aa) , ch-is-sup u u≤x is-sup (fsuc _ fc ) ⟫
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
686 pinit : {y₁ : Ordinal} → odef pchain y₁ → * y ≤ * y₁
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
687 pinit {a} ⟪ aa , ua ⟫ with ua
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
688 ... | ch-init fc = s≤fc y f mf fc
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
689 ... | ch-is-sup u u≤x is-sup fc = ≤-ftrans (<=to≤ zc7) (s≤fc _ f mf fc) where
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
690 zc7 : y <= (ZChain.supf zc) u
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
691 zc7 = ChainP.fcy<sup is-sup (init ay refl)
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
692 pcy : odef pchain y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
693 pcy = ⟪ ay , ch-init (init ay refl) ⟫
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
694
754
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 750
diff changeset
695 supf0 = ZChain.supf zc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 750
diff changeset
696
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
697 sup1 : SUP A (UnionCF A f mf ay supf0 x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
698 sup1 = supP pchain pchain⊆A ptotal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
699 sp1 = & (SUP.sup sup1 )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
700 supf1 : Ordinal → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
701 supf1 z with trio< z px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
702 ... | tri< a ¬b ¬c = ZChain.supf zc z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
703 ... | tri≈ ¬a b ¬c = ZChain.supf zc z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
704 ... | tri> ¬a ¬b c = sp1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
705
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
706 -- if previous chain satisfies maximality, we caan reuse it
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
707 --
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
708 -- supf0 px is sup of UnionCF px , supf0 x is sup of UnionCF x
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
709 no-extension : ¬ sp1 ≡ x → ZChain A f mf ay x
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
710 no-extension ¬sp=x = record { supf = supf1 ; supf-mono = {!!} ; sup = sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
711 ; initial = {!!} ; chain∋init = {!!} ; sup=u = {!!} ; supf-is-sup = {!!} ; csupf = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
712 ; chain⊆A = {!!} ; f-next = {!!} ; f-total = {!!} } where
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
713 UnionCF⊆ : UnionCF A f mf ay supf1 x ⊆' UnionCF A f mf ay supf0 x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
714 UnionCF⊆ ⟪ as , ch-init fc ⟫ = UnionCF⊆ ⟪ as , ch-init fc ⟫
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
715 UnionCF⊆ ⟪ as , ch-is-sup u {z} u≤x record { fcy<sup = f1 ; order = o1 ; supu=u = su=u1 } fc ⟫ with trio< u px
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
716 ... | tri< a ¬b ¬c = ⟪ as , ch-is-sup u {z} u≤x record { fcy<sup = f1 ; order = order0 ; supu=u = {!!} } fc ⟫ where
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
717 order0 : {s z1 : Ordinal} → supf0 s o< supf0 u → FClosure A f (supf0 s) z1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
718 → (z1 ≡ supf0 u) ∨ (z1 << supf0 u)
792
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
719 order0 {s} {z1} ss<su fc with trio< s px | inspect supf1 s
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
720 ... | tri< a ¬b ¬c | record {eq = eq1} = o1 {s} {z1} (subst (λ k → k o< supf0 u) (sym eq1) ss<su )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
721 (subst (λ k → FClosure A f k z1 ) (sym eq1) fc )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
722 ... | tri≈ ¬a b ¬c | record {eq = eq1} = o1 {s} {z1} (subst (λ k → k o< supf0 u) (sym eq1) ss<su )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
723 (subst (λ k → FClosure A f k z1 ) (sym eq1) fc )
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
724 ... | tri> ¬a ¬b c | record {eq = eq1} = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
725 ... | tri≈ ¬a b ¬c = ⟪ as , ch-is-sup u {z} u≤x record { fcy<sup = f1 ; order = order0 ; supu=u = {!!}} fc ⟫ where
792
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
726 order0 : {s z1 : Ordinal} → supf0 s o< supf0 u → FClosure A f (supf0 s) z1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
727 → (z1 ≡ supf0 u) ∨ (z1 << supf0 u)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
728 order0 {s} {z1} ss<su fc with trio< s px | inspect supf1 s
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
729 ... | tri< a ¬b ¬c | record {eq = eq1} = o1 {s} {z1} (subst (λ k → k o< supf0 u) (sym eq1) ss<su )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
730 (subst (λ k → FClosure A f k z1 ) (sym eq1) fc )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
731 ... | tri≈ ¬a b ¬c | record {eq = eq1} = o1 {s} {z1} (subst (λ k → k o< supf0 u) (sym eq1) ss<su )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 791
diff changeset
732 (subst (λ k → FClosure A f k z1 ) (sym eq1) fc )
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
733 ... | tri> ¬a ¬b px<s | record {eq = eq1} = ⊥-elim ( ¬sp=x (subst (λ k → sp1 ≡ k ) u=x {!!} )) where
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
734 s≤u : s o≤ u
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
735 s≤u = {!!}
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
736 u=x : u ≡ x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
737 u=x with trio< u x
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
738 ... | tri< a ¬b ¬c = {!!}
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
739 ... | tri≈ ¬a b ¬c = b
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
740 ... | tri> ¬a ¬b c = {!!}
799
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 798
diff changeset
741 ... | tri> ¬a ¬b c = ⊥-elim ( ¬sp=x (subst (λ k → sp1 ≡ k ) u=x su=u1 )) where
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
742 u=x : u ≡ x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
743 u=x with trio< u x
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
744 ... | tri< a ¬b ¬c = {!!}
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
745 ... | tri≈ ¬a b ¬c = b
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
746 ... | tri> ¬a ¬b c = {!!}
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
747 sup : {z : Ordinal} → z o≤ x → SUP A (UnionCF A f mf ay supf1 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
748 sup {z} z≤x with trio< z px
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
749 ... | tri< a ¬b ¬c = SUP⊆ {!!} (ZChain.sup zc (o<→≤ a))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
750 ... | tri≈ ¬a b ¬c = SUP⊆ {!!} (ZChain.sup zc (o≤-refl0 b))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
751 ... | tri> ¬a ¬b c = SUP⊆ {!!} sup1
709
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 708
diff changeset
752
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
753 zc4 : ZChain A f mf ay x
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
754 zc4 with ODC.∋-p O A (* x)
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
755 ... | no noax = no-extension {!!} -- ¬ A ∋ p, just skip
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
756 ... | yes ax with ODC.p∨¬p O ( HasPrev A (ZChain.chain zc ) ax f )
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
757 -- we have to check adding x preserve is-max ZChain A y f mf x
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
758 ... | case1 pr = no-extension {!!} -- we have previous A ∋ z < x , f z ≡ x, so chain ∋ f z ≡ x because of f-next
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
759 ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A (ZChain.chain zc ) ax )
682
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
760 ... | case1 is-sup = -- x is a sup of zc
786
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 785
diff changeset
761 record { supf = psupf1 ; chain⊆A = {!!} ; f-next = {!!} ; f-total = {!!} ; csupf = {!!} ; sup=u = {!!}
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
762 ; supf-mono = {!!} ; initial = {!!} ; chain∋init = {!!} ; sup = {!!} ; supf-is-sup = {!!} ; supf-mono = {!!} } where
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
763 supx : SUP A (UnionCF A f mf ay supf0 x)
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
764 supx = record { sup = * x ; A∋maximal = subst (λ k → odef A k ) {!!} ax ; x<sup = {!!} }
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
765 spx = & (SUP.sup supx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
766 x=spx : x ≡ spx
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
767 x=spx = {!!}
750
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 749
diff changeset
768 psupf1 : Ordinal → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 749
diff changeset
769 psupf1 z with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 749
diff changeset
770 ... | tri< a ¬b ¬c = ZChain.supf zc z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 749
diff changeset
771 ... | tri≈ ¬a b ¬c = x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 749
diff changeset
772 ... | tri> ¬a ¬b c = x
793
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 792
diff changeset
773
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
774 ... | case2 ¬x=sup = no-extension {!!} -- px is not f y' nor sup of former ZChain from y -- no extention
758
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
775
728
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 727
diff changeset
776 ... | no lim = zc5 where
726
b2e2cd12e38f psupf-mono and is-max conflict
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 725
diff changeset
777
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
778 pzc : (z : Ordinal) → z o< x → ZChain A f mf ay z
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
779 pzc z z<x = prev z z<x
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
780 ysp = & (SUP.sup (ysup f mf ay))
726
b2e2cd12e38f psupf-mono and is-max conflict
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 725
diff changeset
781
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
782 psupf0 : (z : Ordinal) → Ordinal
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
783 psupf0 z with trio< z x
755
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
784 ... | tri< a ¬b ¬c = ZChain.supf (pzc (osuc z) (ob<x lim a)) z
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
785 ... | tri≈ ¬a b ¬c = ysp
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
786 ... | tri> ¬a ¬b c = ysp
755
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
787
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
788 pchain0 : HOD
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
789 pchain0 = UnionCF A f mf ay psupf0 x
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
790
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
791 ptotal0 : IsTotalOrderSet pchain0
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
792 ptotal0 {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
793 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
794 uz01 = chain-total A f mf ay psupf0 ( (proj2 ca)) ( (proj2 cb))
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
795
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
796 usup : SUP A pchain0
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
797 usup = supP pchain0 (λ lt → proj1 lt) ptotal0
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
798 spu = & (SUP.sup usup)
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
799
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
800 supf1 : Ordinal → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
801 supf1 z with trio< z x
755
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
802 ... | tri< a ¬b ¬c = ZChain.supf (pzc (osuc z) (ob<x lim a)) z
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
803 ... | tri≈ ¬a b ¬c = spu
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
804 ... | tri> ¬a ¬b c = spu
755
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
805
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
806 pchain : HOD
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
807 pchain = UnionCF A f mf ay supf1 x
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
808
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
809 pchain⊆A : {y : Ordinal} → odef pchain y → odef A y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
810 pchain⊆A {y} ny = proj1 ny
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
811 pnext : {a : Ordinal} → odef pchain a → odef pchain (f a)
750
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 749
diff changeset
812 pnext {a} ⟪ aa , ch-init fc ⟫ = ⟪ proj2 ( mf a aa ) , ch-init (fsuc _ fc) ⟫
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
813 pnext {a} ⟪ aa , ch-is-sup u u≤x is-sup fc ⟫ = ⟪ proj2 ( mf a aa ) , ch-is-sup u u≤x is-sup (fsuc _ fc) ⟫
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
814 pinit : {y₁ : Ordinal} → odef pchain y₁ → * y ≤ * y₁
748
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
815 pinit {a} ⟪ aa , ua ⟫ with ua
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 747
diff changeset
816 ... | ch-init fc = s≤fc y f mf fc
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
817 ... | ch-is-sup u u≤x is-sup fc = ≤-ftrans (<=to≤ zc7) (s≤fc _ f mf fc) where
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
818 zc7 : y <= supf1 _
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
819 zc7 = ChainP.fcy<sup is-sup (init ay refl)
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
820 pcy : odef pchain y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
821 pcy = ⟪ ay , ch-init (init ay refl) ⟫
755
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
822 ptotal : IsTotalOrderSet pchain
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
823 ptotal {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
824 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
825 uz01 = chain-total A f mf ay supf1 ( (proj2 ca)) ( (proj2 cb))
754
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 750
diff changeset
826
758
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
827 is-max-hp : (supf : Ordinal → Ordinal) (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a →
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
828 b o< x → (ab : odef A b) →
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
829 HasPrev A (UnionCF A f mf ay supf x) ab f →
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
830 * a < * b → odef (UnionCF A f mf ay supf x) b
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
831 is-max-hp supf x {a} {b} ua b<x ab has-prev a<b with HasPrev.ay has-prev
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
832 ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫
791
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 790
diff changeset
833 ... | ⟪ ab0 , ch-is-sup u u≤x is-sup fc ⟫ = ⟪ ab ,
758
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
834 subst (λ k → UChain A f mf ay supf x k )
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
835 (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u≤x is-sup (fsuc _ fc)) ⟫
758
a2947dfff80d is-max on first transfinite induction is not good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 757
diff changeset
836
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
837 no-extension : ¬ spu ≡ x → ZChain A f mf ay x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
838 no-extension ¬sp=x = record { initial = pinit ; chain∋init = pcy ; supf = supf1 ; sup=u = {!!}
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
839 ; sup = sup ; supf-is-sup = sis
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
840 ; csupf = csupf ; chain⊆A = pchain⊆A ; f-next = pnext ; f-total = ptotal ; supf-mono = {!!} } where
795
408e7e8a3797 csupf depends on order cyclicly
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 794
diff changeset
841 supfu : {u : Ordinal } → ( a : u o< x ) → (z : Ordinal) → Ordinal
408e7e8a3797 csupf depends on order cyclicly
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 794
diff changeset
842 supfu {u} a z = ZChain.supf (pzc (osuc u) (ob<x lim a)) z
408e7e8a3797 csupf depends on order cyclicly
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 794
diff changeset
843 UnionCF⊆ : {u : Ordinal} → (a : u o< x ) → UnionCF A f mf ay supf1 x ⊆' UnionCF A f mf ay (supfu a) x
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
844 UnionCF⊆ = {!!}
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
845 sup : {z : Ordinal} → z o≤ x → SUP A (UnionCF A f mf ay supf1 z)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
846 sup {z} z≤x with trio< z x
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
847 ... | tri< a ¬b ¬c = SUP⊆ {!!} (ZChain.sup (pzc z a) o≤-refl )
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
848 ... | tri≈ ¬a b ¬c = SUP⊆ {!!} usup
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
849 ... | tri> ¬a ¬b c = SUP⊆ {!!} usup
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
850 sis : {z : Ordinal} (x≤z : z o≤ x) → supf1 z ≡ & (SUP.sup (sup x≤z))
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
851 sis {z} z≤x with trio< z x
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
852 ... | tri< a ¬b ¬c = ? where
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
853 zc8 = ZChain.supf-is-sup (pzc z a) o≤-refl
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
854 ... | tri≈ ¬a b ¬c = refl
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
855 ... | tri> ¬a ¬b c with osuc-≡< z≤x
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
856 ... | case1 eq = ⊥-elim ( ¬b eq )
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
857 ... | case2 lt = ⊥-elim ( ¬a lt )
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
858 sup=u : {b : Ordinal} (ab : odef A b) → b o< x → IsSup A (UnionCF A f mf ay supf1 (osuc b)) ab → supf1 b ≡ b
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
859 sup=u {b} ab b<x is-sup with trio< b x
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
860 ... | tri< a ¬b ¬c = ZChain.sup=u (pzc (osuc b) (ob<x lim a)) ab <-osuc record { x<sup = ? }
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
861 ... | tri≈ ¬a b ¬c = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
862 ... | tri> ¬a ¬b c = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
863 csupf : {z : Ordinal} → z o≤ x → odef (UnionCF A f mf ay supf1 z) (supf1 z)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
864 csupf {z} z<x with trio< z x
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
865 ... | tri< a ¬b ¬c = zc9 where
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
866 zc9 : odef (UnionCF A f mf ay supf1 z) (ZChain.supf (pzc (osuc z) (ob<x lim a)) z)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
867 zc9 = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
868 zc8 : odef (UnionCF A f mf ay (supfu a) z) (ZChain.supf (pzc (osuc z) (ob<x lim a)) z)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
869 zc8 = ZChain.csupf (pzc (osuc z) (ob<x lim a)) (o<→≤ <-osuc )
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
870 ... | tri≈ ¬a b ¬c = ? -- ⊥-elim (¬a z<x)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
871 ... | tri> ¬a ¬b c = ? -- ⊥-elim (¬a z<x)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
872 supf-mono : {a b : Ordinal} → a o< b → supf1 a o≤ supf1 b
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
873 supf-mono {a0} {b0} a<b = zc10 where
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
874 -- x o≤ a → supf1 a ≡ supf1 b ≡ spu
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
875 -- x o≤ b → supf1 b ≡ spu
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
876 -- a o< x → b o≤ x → supf1 (supf1 a) ≡ supf1 a
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
877 -- supf1 (supf1 b) ≡ supf1 b
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
878 usa : odef (UnionCF A f mf ay (supfu ?) (osuc a0)) (supf1 a0)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
879 usa = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
880 usb : odef (UnionCF A f mf ay (supfu ?) (osuc b0)) (supf1 b0)
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
881 usb = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
882 zc10 : supf1 a0 o≤ supf1 b0
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
883 zc10 with trio< a0 x | trio< b0 x
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
884 ... | tri< a ¬b ¬c | tri< a' ¬b' ¬c' = ? where
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
885 zc11 = ZChain.supf-mono (pzc (osuc a0) (ob<x lim a)) a<b
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
886 zc12 = ZChain.supf-mono (pzc (osuc b0) (ob<x lim a')) a<b
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
887 ... | tri< a ¬b ¬c | tri≈ ¬a b ¬c' = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
888 ... | tri< a ¬b ¬c | tri> ¬a ¬b' c = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
889 ... | tri≈ ¬a b ¬c | tri< a ¬b ¬c' = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
890 ... | tri≈ ¬a b ¬c | tri≈ ¬a' b' ¬c' = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
891 ... | tri≈ ¬a b ¬c | tri> ¬a' ¬b c = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
892 ... | tri> ¬a ¬b c | _ = ?
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
893
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
894 zc5 : ZChain A f mf ay x
697
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 696
diff changeset
895 zc5 with ODC.∋-p O A (* x)
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
896 ... | no noax = no-extension {!!} -- ¬ A ∋ p, just skip
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
897 ... | yes ax with ODC.p∨¬p O ( HasPrev A pchain ax f )
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
898 -- we have to check adding x preserve is-max ZChain A y f mf x
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
899 ... | case1 pr = no-extension {!!}
704
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 703
diff changeset
900 ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A pchain ax )
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
901 ... | case1 is-sup = record { initial = {!!} ; chain∋init = {!!} ; supf = supf1 ; sup=u = {!!}
796
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
902 ; sup = {!!} ; supf-is-sup = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
903 ; chain⊆A = {!!} ; f-next = {!!} ; f-total = {!!} ; csupf = {!!} ; supf-mono = {!!} } where -- x is a sup of (zc ?)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 795
diff changeset
904 ... | case2 ¬x=sup = no-extension {!!} -- x is not f y' nor sup of former ZChain from y -- no extention
553
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 552
diff changeset
905
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
906 SZ : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → ZChain A f mf ay (& A)
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
907 SZ f mf {y} ay = TransFinite {λ z → ZChain A f mf ay z } (λ x → ind f mf ay x ) (& A)
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
908
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
909 zorn00 : Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
910 zorn00 with is-o∅ ( & HasMaximal ) -- we have no Level (suc n) LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
911 ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; A∋maximal = zorn01 ; ¬maximal<x = zorn02 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
912 -- yes we have the maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
913 zorn03 : odef HasMaximal ( & ( ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ) )
606
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
914 zorn03 = ODC.x∋minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) -- Axiom of choice
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
915 zorn01 : A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
916 zorn01 = proj1 zorn03
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
917 zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
918 zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x )
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
919 ... | yes ¬Maximal = ⊥-elim ( z04 nmx zorn04 total ) where
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
920 -- if we have no maximal, make ZChain, which contradict SUP condition
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
921 nmx : ¬ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
922 nmx mx = ∅< {HasMaximal} zc5 ( ≡o∅→=od∅ ¬Maximal ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
923 zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) → odef A y → ¬ (* (& (Maximal.maximal mx)) < * y))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
924 zc5 = ⟪ Maximal.A∋maximal mx , (λ y ay mx<y → Maximal.¬maximal<x mx (subst (λ k → odef A k ) (sym &iso) ay) (subst (λ k → k < * y) *iso mx<y) ) ⟫
703
0278f0d151f2 one pass
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 698
diff changeset
925 zorn04 : ZChain A (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A)
653
4186c0331abb sind again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 638
diff changeset
926 zorn04 = SZ (cf nmx) (cf-is-≤-monotonic nmx) (subst (λ k → odef A k ) &iso as )
634
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 633
diff changeset
927 total : IsTotalOrderSet (ZChain.chain zorn04)
654
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 653
diff changeset
928 total {a} {b} = zorn06 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 653
diff changeset
929 zorn06 : odef (ZChain.chain zorn04) (& a) → odef (ZChain.chain zorn04) (& b) → Tri (a < b) (a ≡ b) (b < a)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 653
diff changeset
930 zorn06 = ZChain.f-total (SZ (cf nmx) (cf-is-≤-monotonic nmx) (subst (λ k → odef A k ) &iso as) )
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
931
516
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
932 -- usage (see filter.agda )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
933 --
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
934 -- _⊆'_ : ( A B : HOD ) → Set n
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
935 -- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x
482
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 481
diff changeset
936
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
937 -- MaximumSubset : {L P : HOD}
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
938 -- → o∅ o< & L → o∅ o< & P → P ⊆ L
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
939 -- → IsPartialOrderSet P _⊆'_
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
940 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ )
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
941 -- → Maximal P (_⊆'_)
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
942 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆'_} 0<P PO SP