comparison src/zorn.agda @ 966:39c680188738

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 05 Nov 2022 13:21:42 +0900
parents 1c1c6a6ed4fa
children cd0ef83189c5
comparison
equal deleted inserted replaced
965:1c1c6a6ed4fa 966:39c680188738
1 {-# OPTIONS --allow-unsolved-metas #-} 1 {-# OPTIONS --allow-unsolved-metas #-}
2 open import Level hiding ( suc ; zero ) 2 open import Level hiding ( suc ; zero )
3 open import Ordinals 3 open import Ordinals
4 open import Relation.Binary 4 open import Relation.Binary
5 open import Relation.Binary.Core 5 open import Relation.Binary.Core
6 open import Relation.Binary.PropositionalEquality 6 open import Relation.Binary.PropositionalEquality
7 import OD 7 import OD
8 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where 8 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where
9 9
10 -- 10 --
11 -- Zorn-lemma : { A : HOD } 11 -- Zorn-lemma : { A : HOD }
12 -- → o∅ o< & A 12 -- → o∅ o< & A
13 -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition 13 -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
14 -- → Maximal A 14 -- → Maximal A
15 -- 15 --
16 16
17 open import zf 17 open import zf
18 open import logic 18 open import logic
19 -- open import partfunc {n} O 19 -- open import partfunc {n} O
20 20
21 open import Relation.Nullary 21 open import Relation.Nullary
22 open import Data.Empty 22 open import Data.Empty
23 import BAlgbra 23 import BAlgbra
24 24
25 open import Data.Nat hiding ( _<_ ; _≤_ ) 25 open import Data.Nat hiding ( _<_ ; _≤_ )
26 open import Data.Nat.Properties 26 open import Data.Nat.Properties
27 open import nat 27 open import nat
28 28
29 29
30 open inOrdinal O 30 open inOrdinal O
31 open OD O 31 open OD O
32 open OD.OD 32 open OD.OD
50 50
51 -- 51 --
52 -- Partial Order on HOD ( possibly limited in A ) 52 -- Partial Order on HOD ( possibly limited in A )
53 -- 53 --
54 54
55 _<<_ : (x y : Ordinal ) → Set n 55 _<<_ : (x y : Ordinal ) → Set n
56 x << y = * x < * y 56 x << y = * x < * y
57 57
58 _<=_ : (x y : Ordinal ) → Set n -- we can't use * x ≡ * y, it is Set (Level.suc n). Level (suc n) troubles Chain 58 _<=_ : (x y : Ordinal ) → Set n -- we can't use * x ≡ * y, it is Set (Level.suc n). Level (suc n) troubles Chain
59 x <= y = (x ≡ y ) ∨ ( * x < * y ) 59 x <= y = (x ≡ y ) ∨ ( * x < * y )
60 60
61 POO : IsStrictPartialOrder _≡_ _<<_ 61 POO : IsStrictPartialOrder _≡_ _<<_
62 POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans } 62 POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans }
63 ; trans = IsStrictPartialOrder.trans PO 63 ; trans = IsStrictPartialOrder.trans PO
64 ; irrefl = λ x=y x<y → IsStrictPartialOrder.irrefl PO (cong (*) x=y) x<y 64 ; irrefl = λ x=y x<y → IsStrictPartialOrder.irrefl PO (cong (*) x=y) x<y
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 } } 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 } }
66 66
67 _≤_ : (x y : HOD) → Set (Level.suc n) 67 _≤_ : (x y : HOD) → Set (Level.suc n)
68 x ≤ y = ( x ≡ y ) ∨ ( x < y ) 68 x ≤ y = ( x ≡ y ) ∨ ( x < y )
69 69
70 ≤-ftrans : {x y z : HOD} → x ≤ y → y ≤ z → x ≤ z 70 ≤-ftrans : {x y z : HOD} → x ≤ y → y ≤ z → x ≤ z
71 ≤-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl 71 ≤-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
72 ≤-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z 72 ≤-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
73 ≤-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y 73 ≤-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
74 ≤-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z ) 74 ≤-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
75 75
76 <=-trans : {x y z : Ordinal } → x <= y → y <= z → x <= z 76 <=-trans : {x y z : Ordinal } → x <= y → y <= z → x <= z
77 <=-trans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl 77 <=-trans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
78 <=-trans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z 78 <=-trans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
79 <=-trans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y 79 <=-trans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
80 <=-trans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z ) 80 <=-trans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
81 81
82 ftrans<=-< : {x y z : Ordinal } → x <= y → y << z → x << z 82 ftrans<=-< : {x y z : Ordinal } → x <= y → y << z → x << z
83 ftrans<=-< {x} {y} {z} (case1 eq) y<z = subst (λ k → k < * z) (sym (cong (*) eq)) y<z 83 ftrans<=-< {x} {y} {z} (case1 eq) y<z = subst (λ k → k < * z) (sym (cong (*) eq)) y<z
84 ftrans<=-< {x} {y} {z} (case2 lt) y<z = IsStrictPartialOrder.trans PO lt y<z 84 ftrans<=-< {x} {y} {z} (case2 lt) y<z = IsStrictPartialOrder.trans PO lt y<z
85 85
86 <=to≤ : {x y : Ordinal } → x <= y → * x ≤ * y 86 <=to≤ : {x y : Ordinal } → x <= y → * x ≤ * y
87 <=to≤ (case1 eq) = case1 (cong (*) eq) 87 <=to≤ (case1 eq) = case1 (cong (*) eq)
88 <=to≤ (case2 lt) = case2 lt 88 <=to≤ (case2 lt) = case2 lt
89 89
90 ≤to<= : {x y : Ordinal } → * x ≤ * y → x <= y 90 ≤to<= : {x y : Ordinal } → * x ≤ * y → x <= y
91 ≤to<= (case1 eq) = case1 ( subst₂ (λ j k → j ≡ k ) &iso &iso (cong (&) eq) ) 91 ≤to<= (case1 eq) = case1 ( subst₂ (λ j k → j ≡ k ) &iso &iso (cong (&) eq) )
92 ≤to<= (case2 lt) = case2 lt 92 ≤to<= (case2 lt) = case2 lt
93 93
94 <-irr : {a b : HOD} → (a ≡ b ) ∨ (a < b ) → b < a → ⊥ 94 <-irr : {a b : HOD} → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
95 <-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (sym a=b) b<a 95 <-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (sym a=b) b<a
99 ptrans = IsStrictPartialOrder.trans PO 99 ptrans = IsStrictPartialOrder.trans PO
100 100
101 open _==_ 101 open _==_
102 open _⊆_ 102 open _⊆_
103 103
104 -- <-TransFinite : {A x : HOD} → {P : HOD → Set n} → x ∈ A 104 -- <-TransFinite : {A x : HOD} → {P : HOD → Set n} → x ∈ A
105 -- → ({x : HOD} → A ∋ x → ({y : HOD} → A ∋ y → y < x → P y ) → P x) → P x 105 -- → ({x : HOD} → A ∋ x → ({y : HOD} → A ∋ y → y < x → P y ) → P x) → P x
106 -- <-TransFinite = ? 106 -- <-TransFinite = ?
107 107
108 -- 108 --
109 -- Closure of ≤-monotonic function f has total order 109 -- Closure of ≤-monotonic function f has total order
120 A∋fc {A} s f mf (init as refl ) = as 120 A∋fc {A} s f mf (init as refl ) = as
121 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) ) 121 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) )
122 122
123 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 123 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
124 A∋fcs {A} s f mf (init as refl) = as 124 A∋fcs {A} s f mf (init as refl) = as
125 A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy 125 A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy
126 126
127 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → * s ≤ * y 127 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → * s ≤ * y
128 s≤fc {A} s {.s} f mf (init x refl ) = case1 refl 128 s≤fc {A} s {.s} f mf (init x refl ) = case1 refl
129 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) ) 129 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) )
130 ... | case1 x=fx = subst (λ k → * s ≤ * k ) (*≡*→≡ x=fx) ( s≤fc {A} s f mf fcy ) 130 ... | case1 x=fx = subst (λ k → * s ≤ * k ) (*≡*→≡ x=fx) ( s≤fc {A} s f mf fcy )
131 ... | case2 x<fx with s≤fc {A} s f mf fcy 131 ... | case2 x<fx with s≤fc {A} s f mf fcy
132 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym s≡x) refl x<fx ) 132 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym s≡x) refl x<fx )
133 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx ) 133 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx )
134 134
135 fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ 135 fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ
136 fcn s mf (init as refl) = zero 136 fcn s mf (init as refl) = zero
137 fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p)) 137 fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p))
138 ... | case1 eq = fcn s mf p 138 ... | case1 eq = fcn s mf p
139 ... | case2 y<fy = suc (fcn s mf p ) 139 ... | case2 y<fy = suc (fcn s mf p )
140 140
141 fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) 141 fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
142 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx ≡ fcn s mf cy → * x ≡ * y 142 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx ≡ fcn s mf cy → * x ≡ * y
143 fcn-inject {A} s {x} {y} {f} mf cx cy eq = fc00 (fcn s mf cx) (fcn s mf cy) eq cx cy refl refl where 143 fcn-inject {A} s {x} {y} {f} mf cx cy eq = fc00 (fcn s mf cx) (fcn s mf cy) eq cx cy refl refl where
144 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq ) 144 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq )
145 fc06 {x} {y} refl {j} not = fc08 not where 145 fc06 {x} {y} refl {j} not = fc08 not where
146 fc08 : {j : ℕ} → ¬ suc j ≡ 0 146 fc08 : {j : ℕ} → ¬ suc j ≡ 0
147 fc08 () 147 fc08 ()
148 fc07 : {x : Ordinal } (cx : FClosure A f s x ) → 0 ≡ fcn s mf cx → * s ≡ * x 148 fc07 : {x : Ordinal } (cx : FClosure A f s x ) → 0 ≡ fcn s mf cx → * s ≡ * x
149 fc07 {x} (init as refl) eq = refl 149 fc07 {x} (init as refl) eq = refl
150 fc07 {.(f x)} (fsuc x cx) eq with proj1 (mf x (A∋fc s f mf cx ) ) 150 fc07 {.(f x)} (fsuc x cx) eq with proj1 (mf x (A∋fc s f mf cx ) )
151 ... | case1 x=fx = subst (λ k → * s ≡ k ) x=fx ( fc07 cx eq ) 151 ... | case1 x=fx = subst (λ k → * s ≡ k ) x=fx ( fc07 cx eq )
172 fc04 = fc00 i j (cong pred i=j) cx1 cy (cong pred i=x1) (cong pred j=y) 172 fc04 = fc00 i j (cong pred i=j) cx1 cy (cong pred i=x1) (cong pred j=y)
173 ... | case2 x<fx | case1 y=fy = subst (λ k → * (f x) ≡ k ) y=fy (fc03 y cy j=y) where 173 ... | case2 x<fx | case1 y=fy = subst (λ k → * (f x) ≡ k ) y=fy (fc03 y cy j=y) where
174 fc03 : (y1 : Ordinal) → (cy1 : FClosure A f s y1 ) → suc j ≡ fcn s mf cy1 → * (f x) ≡ * y1 174 fc03 : (y1 : Ordinal) → (cy1 : FClosure A f s y1 ) → suc j ≡ fcn s mf cy1 → * (f x) ≡ * y1
175 fc03 y1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x) 175 fc03 y1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x)
176 fc03 .(f y1) (fsuc y1 cy1) j=y1 with proj1 (mf y1 (A∋fc s f mf cy1 ) ) 176 fc03 .(f y1) (fsuc y1 cy1) j=y1 with proj1 (mf y1 (A∋fc s f mf cy1 ) )
177 ... | case1 eq = trans ( fc03 y1 cy1 j=y1 ) eq 177 ... | case1 eq = trans ( fc03 y1 cy1 j=y1 ) eq
178 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where 178 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where
179 fc05 : * x ≡ * y1 179 fc05 : * x ≡ * y1
180 fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1) 180 fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1)
181 ... | case2 x₁ | case2 x₂ = subst₂ (λ j k → * (f j) ≡ * (f k) ) &iso &iso (cong (λ k → * (f (& k))) (fc00 i j (cong pred i=j) cx cy (cong pred i=x) (cong pred j=y))) 181 ... | case2 x₁ | case2 x₂ = subst₂ (λ j k → * (f j) ≡ * (f k) ) &iso &iso (cong (λ k → * (f (& k))) (fc00 i j (cong pred i=j) cx cy (cong pred i=x) (cong pred j=y)))
182 182
184 fcn-< : {A : HOD} (s : Ordinal ) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) 184 fcn-< : {A : HOD} (s : Ordinal ) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
185 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx Data.Nat.< fcn s mf cy → * x < * y 185 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx Data.Nat.< fcn s mf cy → * x < * y
186 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where 186 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where
187 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq ) 187 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq )
188 fc06 {x} {y} refl {j} not = fc08 not where 188 fc06 {x} {y} refl {j} not = fc08 not where
189 fc08 : {j : ℕ} → ¬ suc j ≡ 0 189 fc08 : {j : ℕ} → ¬ suc j ≡ 0
190 fc08 () 190 fc08 ()
191 fc01 : (i : ℕ ) → {y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → (i ≡ fcn s mf cy ) → fcn s mf cx Data.Nat.< i → * x < * y 191 fc01 : (i : ℕ ) → {y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → (i ≡ fcn s mf cy ) → fcn s mf cx Data.Nat.< i → * x < * y
192 fc01 (suc i) cx (init x₁ x₂) x (s≤s x₃) = ⊥-elim (fc06 x₂ x) 192 fc01 (suc i) cx (init x₁ x₂) x (s≤s x₃) = ⊥-elim (fc06 x₂ x)
193 fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) ) 193 fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) )
194 ... | case1 y=fy = subst (λ k → * x < k ) y=fy ( fc01 (suc i) {y1} cx cy i=y (s≤s x<i) ) 194 ... | case1 y=fy = subst (λ k → * x < k ) y=fy ( fc01 (suc i) {y1} cx cy i=y (s≤s x<i) )
195 ... | case2 y<fy with <-cmp (fcn s mf cx ) i 195 ... | case2 y<fy with <-cmp (fcn s mf cx ) i
196 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> x<i c ) 196 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> x<i c )
197 ... | tri≈ ¬a b ¬c = subst (λ k → k < * (f y1) ) (fcn-inject s mf cy cx (sym (trans b (cong pred i=y) ))) y<fy 197 ... | tri≈ ¬a b ¬c = subst (λ k → k < * (f y1) ) (fcn-inject s mf cy cx (sym (trans b (cong pred i=y) ))) y<fy
198 ... | tri< a ¬b ¬c = IsStrictPartialOrder.trans PO fc02 y<fy where 198 ... | tri< a ¬b ¬c = IsStrictPartialOrder.trans PO fc02 y<fy where
199 fc03 : suc i ≡ suc (fcn s mf cy) → i ≡ fcn s mf cy 199 fc03 : suc i ≡ suc (fcn s mf cy) → i ≡ fcn s mf cy
200 fc03 eq = cong pred eq 200 fc03 eq = cong pred eq
201 fc02 : * x < * y1 201 fc02 : * x < * y1
202 fc02 = fc01 i cx cy (fc03 i=y ) a 202 fc02 = fc01 i cx cy (fc03 i=y ) a
203 203
204 204
205 fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) 205 fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
206 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → Tri (* x < * y) (* x ≡ * y) (* y < * x ) 206 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → Tri (* x < * y) (* x ≡ * y) (* y < * x )
207 fcn-cmp {A} s {x} {y} f mf cx cy with <-cmp ( fcn s mf cx ) (fcn s mf cy ) 207 fcn-cmp {A} s {x} {y} f mf cx cy with <-cmp ( fcn s mf cx ) (fcn s mf cy )
208 ... | tri< a ¬b ¬c = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where 208 ... | tri< a ¬b ¬c = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where
209 fc11 : * x < * y 209 fc11 : * x < * y
210 fc11 = fcn-< {A} s {x} {y} {f} mf cx cy a 210 fc11 = fcn-< {A} s {x} {y} {f} mf cx cy a
211 ... | tri≈ ¬a b ¬c = tri≈ (λ lt → <-irr (case1 (sym fc10)) lt) fc10 (λ lt → <-irr (case1 fc10) lt) where 211 ... | tri≈ ¬a b ¬c = tri≈ (λ lt → <-irr (case1 (sym fc10)) lt) fc10 (λ lt → <-irr (case1 fc10) lt) where
212 fc10 : * x ≡ * y 212 fc10 : * x ≡ * y
213 fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b 213 fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b
214 ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where 214 ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where
215 fc12 : * y < * x 215 fc12 : * y < * x
216 fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c 216 fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c
217 217
218 218
219 219
236 record HasPrev (A B : HOD) ( f : Ordinal → Ordinal ) (x : Ordinal ) : Set n where 236 record HasPrev (A B : HOD) ( f : Ordinal → Ordinal ) (x : Ordinal ) : Set n where
237 field 237 field
238 ax : odef A x 238 ax : odef A x
239 y : Ordinal 239 y : Ordinal
240 ay : odef B y 240 ay : odef B y
241 x=fy : x ≡ f y 241 x=fy : x ≡ f y
242 242
243 record IsSUP (A B : HOD) {x : Ordinal } (xa : odef A x) : Set n where 243 record IsSUP (A B : HOD) {x : Ordinal } (xa : odef A x) : Set n where
244 field 244 field
245 x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) 245 x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x )
246 246
247 record IsMinSUP (A B : HOD) ( f : Ordinal → Ordinal ) {x : Ordinal } (xa : odef A x) : Set n where 247 record IsMinSUP (A B : HOD) ( f : Ordinal → Ordinal ) {x : Ordinal } (xa : odef A x) : Set n where
248 field 248 field
249 x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) 249 x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x )
250 minsup : { sup1 : Ordinal } → odef A sup1 250 minsup : { sup1 : Ordinal } → odef A sup1
251 → ( {z : Ordinal } → odef B z → (z ≡ sup1 ) ∨ (z << sup1 )) → x o≤ sup1 251 → ( {z : Ordinal } → odef B z → (z ≡ sup1 ) ∨ (z << sup1 )) → x o≤ sup1
252 not-hp : ¬ ( HasPrev A B f x ) 252 not-hp : ¬ ( HasPrev A B f x )
253 253
254 record SUP ( A B : HOD ) : Set (Level.suc n) where 254 record SUP ( A B : HOD ) : Set (Level.suc n) where
255 field 255 field
259 259
260 -- 260 --
261 -- sup and its fclosure is in a chain HOD 261 -- sup and its fclosure is in a chain HOD
262 -- chain HOD is sorted by sup as Ordinal and <-ordered 262 -- chain HOD is sorted by sup as Ordinal and <-ordered
263 -- whole chain is a union of separated Chain 263 -- whole chain is a union of separated Chain
264 -- minimum index is sup of y not ϕ 264 -- minimum index is sup of y not ϕ
265 -- 265 --
266 266
267 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 267 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
268 field 268 field
269 fcy<sup : {z : Ordinal } → FClosure A f y z → (z ≡ supf u) ∨ ( z << supf u ) 269 fcy<sup : {z : Ordinal } → FClosure A f y z → (z ≡ supf u) ∨ ( z << supf u )
270 order : {s z1 : Ordinal} → (lt : supf s o< supf u ) → FClosure A f (supf s ) z1 → (z1 ≡ supf u ) ∨ ( z1 << supf u ) 270 order : {s z1 : Ordinal} → (lt : supf s o< supf u ) → FClosure A f (supf s ) z1 → (z1 ≡ supf u ) ∨ ( z1 << supf u )
271 supu=u : supf u ≡ u 271 supu=u : supf u ≡ u
272 272
273 data UChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y ) 273 data UChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y )
274 (supf : Ordinal → Ordinal) (x : Ordinal) : (z : Ordinal) → Set n where 274 (supf : Ordinal → Ordinal) (x : Ordinal) : (z : Ordinal) → Set n where
275 ch-is-sup : (u : Ordinal) {z : Ordinal } (u<x : supf u o< supf x) ( is-sup : ChainP A f mf ay supf u ) 275 ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain A f mf ay supf x z
276 ch-is-sup : (u : Ordinal) {z : Ordinal } (u<x : supf u o< supf x) ( is-sup : ChainP A f mf ay supf u )
276 ( fc : FClosure A f (supf u) z ) → UChain A f mf ay supf x z 277 ( fc : FClosure A f (supf u) z ) → UChain A f mf ay supf x z
277 278
278 -- 279 --
279 -- f (f ( ... (sup y))) f (f ( ... (sup z1))) 280 -- f (f ( ... (sup y))) f (f ( ... (sup z1)))
280 -- / | / | 281 -- / | / |
285 286
286 chain-total : (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal ) 287 chain-total : (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
287 {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 ) 288 {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 )
288 chain-total A f mf {y} ay supf {xa} {xb} {a} {b} ca cb = ct-ind xa xb ca cb where 289 chain-total A f mf {y} ay supf {xa} {xb} {a} {b} ca cb = ct-ind xa xb ca cb where
289 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) 290 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)
291 ct-ind xa xb {a} {b} (ch-init fca) (ch-init fcb) = fcn-cmp y f mf fca fcb
292 ct-ind xa xb {a} {b} (ch-init fca) (ch-is-sup ub u<x supb fcb) with ChainP.fcy<sup supb fca
293 ... | case1 eq with s≤fc (supf ub) f mf fcb
294 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
295 ct00 : * a ≡ * b
296 ct00 = trans (cong (*) eq) eq1
297 ... | case2 lt = tri< ct01 (λ eq → <-irr (case1 (sym eq)) ct01) (λ lt → <-irr (case2 ct01) lt) where
298 ct01 : * a < * b
299 ct01 = subst (λ k → * k < * b ) (sym eq) lt
300 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
301 ct00 : * a < * (supf ub)
302 ct00 = lt
303 ct01 : * a < * b
304 ct01 with s≤fc (supf ub) f mf fcb
305 ... | case1 eq = subst (λ k → * a < k ) eq ct00
306 ... | case2 lt = IsStrictPartialOrder.trans POO ct00 lt
307 ct-ind xa xb {a} {b} (ch-is-sup ua u<x supa fca) (ch-init fcb) with ChainP.fcy<sup supa fcb
308 ... | case1 eq with s≤fc (supf ua) f mf fca
309 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
310 ct00 : * a ≡ * b
311 ct00 = sym (trans (cong (*) eq) eq1 )
312 ... | case2 lt = tri> (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01 where
313 ct01 : * b < * a
314 ct01 = subst (λ k → * k < * a ) (sym eq) lt
315 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
316 ct00 : * b < * (supf ua)
317 ct00 = lt
318 ct01 : * b < * a
319 ct01 with s≤fc (supf ua) f mf fca
320 ... | case1 eq = subst (λ k → * b < k ) eq ct00
321 ... | case2 lt = IsStrictPartialOrder.trans POO ct00 lt
290 ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) with trio< ua ub 322 ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) with trio< ua ub
291 ... | tri< a₁ ¬b ¬c with ChainP.order supb (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supa )) (sym (ChainP.supu=u supb )) a₁) fca 323 ... | tri< a₁ ¬b ¬c with ChainP.order supb (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supa )) (sym (ChainP.supu=u supb )) a₁) fca
292 ... | case1 eq with s≤fc (supf ub) f mf fcb 324 ... | case1 eq with s≤fc (supf ub) f mf fcb
293 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where 325 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
294 ct00 : * a ≡ * b 326 ct00 : * a ≡ * b
295 ct00 = trans (cong (*) eq) eq1 327 ct00 = trans (cong (*) eq) eq1
296 ... | case2 lt = tri< ct02 (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt) where 328 ... | case2 lt = tri< ct02 (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt) where
297 ct02 : * a < * b 329 ct02 : * a < * b
298 ct02 = subst (λ k → * k < * b ) (sym eq) lt 330 ct02 = subst (λ k → * k < * b ) (sym eq) lt
299 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 331 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
300 ct03 : * a < * (supf ub) 332 ct03 : * a < * (supf ub)
301 ct03 = lt 333 ct03 = lt
302 ct02 : * a < * b 334 ct02 : * a < * b
303 ct02 with s≤fc (supf ub) f mf fcb 335 ct02 with s≤fc (supf ub) f mf fcb
304 ... | case1 eq = subst (λ k → * a < k ) eq ct03 336 ... | case1 eq = subst (λ k → * a < k ) eq ct03
305 ... | case2 lt = IsStrictPartialOrder.trans POO ct03 lt 337 ... | case2 lt = IsStrictPartialOrder.trans POO ct03 lt
306 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 338 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
307 = fcn-cmp (supf ua) f mf fca (subst (λ k → FClosure A f k b ) (cong supf (sym eq)) fcb ) 339 = fcn-cmp (supf ua) f mf fca (subst (λ k → FClosure A f k b ) (cong supf (sym eq)) fcb )
308 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 (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supb )) (sym (ChainP.supu=u supa )) c) fcb 340 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 (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supb )) (sym (ChainP.supu=u supa )) c) fcb
309 ... | case1 eq with s≤fc (supf ua) f mf fca 341 ... | case1 eq with s≤fc (supf ua) f mf fca
310 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where 342 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
311 ct00 : * a ≡ * b 343 ct00 : * a ≡ * b
312 ct00 = sym (trans (cong (*) eq) eq1) 344 ct00 = sym (trans (cong (*) eq) eq1)
313 ... | case2 lt = tri> (λ lt → <-irr (case2 ct02) lt) (λ eq → <-irr (case1 eq) ct02) ct02 where 345 ... | case2 lt = tri> (λ lt → <-irr (case2 ct02) lt) (λ eq → <-irr (case1 eq) ct02) ct02 where
314 ct02 : * b < * a 346 ct02 : * b < * a
315 ct02 = subst (λ k → * k < * a ) (sym eq) lt 347 ct02 = subst (λ k → * k < * a ) (sym eq) lt
316 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 348 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
317 ct05 : * b < * (supf ua) 349 ct05 : * b < * (supf ua)
318 ct05 = lt 350 ct05 = lt
319 ct04 : * b < * a 351 ct04 : * b < * a
320 ct04 with s≤fc (supf ua) f mf fca 352 ct04 with s≤fc (supf ua) f mf fca
321 ... | case1 eq = subst (λ k → * b < k ) eq ct05 353 ... | case1 eq = subst (λ k → * b < k ) eq ct05
322 ... | case2 lt = IsStrictPartialOrder.trans POO ct05 lt 354 ... | case2 lt = IsStrictPartialOrder.trans POO ct05 lt
323 355
324 ∈∧P→o< : {A : HOD } {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A 356 ∈∧P→o< : {A : HOD } {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A
325 ∈∧P→o< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p ))) 357 ∈∧P→o< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
326 358
327 -- Union of supf z which o< x 359 -- Union of supf z which o< x
328 -- 360 --
329 UnionCF : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y ) 361 UnionCF : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y )
330 ( supf : Ordinal → Ordinal ) ( x : Ordinal ) → HOD 362 ( supf : Ordinal → Ordinal ) ( x : Ordinal ) → HOD
331 UnionCF A f mf ay supf x 363 UnionCF A f mf ay supf x
332 = record { od = record { def = λ z → odef A z ∧ UChain A f mf ay supf x z } ; odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy } 364 = record { od = record { def = λ z → odef A z ∧ UChain A f mf ay supf x z } ; odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy }
333 365
334 supf-inject0 : {x y : Ordinal } {supf : Ordinal → Ordinal } → (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y ) 366 supf-inject0 : {x y : Ordinal } {supf : Ordinal → Ordinal } → (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y )
335 → supf x o< supf y → x o< y 367 → supf x o< supf y → x o< y
336 supf-inject0 {x} {y} {supf} supf-mono sx<sy with trio< x y 368 supf-inject0 {x} {y} {supf} supf-mono sx<sy with trio< x y
337 ... | tri< a ¬b ¬c = a 369 ... | tri< a ¬b ¬c = a
338 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy ) 370 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy )
339 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) ) 371 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) )
340 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy ) 372 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy )
342 374
343 record MinSUP ( A B : HOD ) : Set n where 375 record MinSUP ( A B : HOD ) : Set n where
344 field 376 field
345 sup : Ordinal 377 sup : Ordinal
346 asm : odef A sup 378 asm : odef A sup
347 x≤sup : {x : Ordinal } → odef B x → (x ≡ sup ) ∨ (x << sup ) 379 x≤sup : {x : Ordinal } → odef B x → (x ≡ sup ) ∨ (x << sup )
348 minsup : { sup1 : Ordinal } → odef A sup1 380 minsup : { sup1 : Ordinal } → odef A sup1
349 → ( {x : Ordinal } → odef B x → (x ≡ sup1 ) ∨ (x << sup1 )) → sup o≤ sup1 381 → ( {x : Ordinal } → odef B x → (x ≡ sup1 ) ∨ (x << sup1 )) → sup o≤ sup1
350 382
351 z09 : {b : Ordinal } { A : HOD } → odef A b → b o< & A 383 z09 : {b : Ordinal } { A : HOD } → odef A b → b o< & A
352 z09 {b} {A} ab = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab)) 384 z09 {b} {A} ab = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
353 385
354 M→S : { A : HOD } { f : Ordinal → Ordinal } {mf : ≤-monotonic-f A f} {y : Ordinal} {ay : odef A y} { x : Ordinal } 386 M→S : { A : HOD } { f : Ordinal → Ordinal } {mf : ≤-monotonic-f A f} {y : Ordinal} {ay : odef A y} { x : Ordinal }
355 → (supf : Ordinal → Ordinal ) 387 → (supf : Ordinal → Ordinal )
356 → MinSUP A (UnionCF A f mf ay supf x) 388 → MinSUP A (UnionCF A f mf ay supf x)
357 → SUP A (UnionCF A f mf ay supf x) 389 → SUP A (UnionCF A f mf ay supf x)
358 M→S {A} {f} {mf} {y} {ay} {x} supf ms = record { sup = * (MinSUP.sup ms) 390 M→S {A} {f} {mf} {y} {ay} {x} supf ms = record { sup = * (MinSUP.sup ms)
359 ; as = subst (λ k → odef A k) (sym &iso) (MinSUP.asm ms) ; x≤sup = ms00 } where 391 ; as = subst (λ k → odef A k) (sym &iso) (MinSUP.asm ms) ; x≤sup = ms00 } where
360 msup = MinSUP.sup ms 392 msup = MinSUP.sup ms
361 ms00 : {z : HOD} → UnionCF A f mf ay supf x ∋ z → (z ≡ * msup) ∨ (z < * msup) 393 ms00 : {z : HOD} → UnionCF A f mf ay supf x ∋ z → (z ≡ * msup) ∨ (z < * msup)
362 ms00 {z} uz with MinSUP.x≤sup ms uz 394 ms00 {z} uz with MinSUP.x≤sup ms uz
363 ... | case1 eq = case1 (subst (λ k → k ≡ _) *iso ( cong (*) eq)) 395 ... | case1 eq = case1 (subst (λ k → k ≡ _) *iso ( cong (*) eq))
364 ... | case2 lt = case2 (subst₂ (λ j k → j < k ) *iso refl lt ) 396 ... | case2 lt = case2 (subst₂ (λ j k → j < k ) *iso refl lt )
365 397
366 398
367 chain-mono : {A : HOD} ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal ) 399 chain-mono : {A : HOD} ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
368 (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y ) {a b c : Ordinal} → a o≤ b 400 (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y ) {a b c : Ordinal} → a o≤ b
369 → odef (UnionCF A f mf ay supf a) c → odef (UnionCF A f mf ay supf b) c 401 → odef (UnionCF A f mf ay supf a) c → odef (UnionCF A f mf ay supf b) c
402 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ ua , ch-init fc ⟫ =
403 ⟪ ua , ch-init fc ⟫
370 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ uaa , ch-is-sup ua ua<x is-sup fc ⟫ = 404 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ uaa , ch-is-sup ua ua<x is-sup fc ⟫ =
371 ⟪ uaa , ch-is-sup ua (ordtrans<-≤ ua<x (supf-mono a≤b ) ) is-sup fc ⟫ 405 ⟪ uaa , ch-is-sup ua (ordtrans<-≤ ua<x (supf-mono a≤b ) ) is-sup fc ⟫
372 406
373 record ZChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) 407 record ZChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)
374 {y : Ordinal} (ay : odef A y) ( z : Ordinal ) : Set (Level.suc n) where 408 {y : Ordinal} (ay : odef A y) ( z : Ordinal ) : Set (Level.suc n) where
375 field 409 field
376 supf : Ordinal → Ordinal 410 supf : Ordinal → Ordinal
377 sup=u : {b : Ordinal} → (ab : odef A b) → b o≤ z 411 sup=u : {b : Ordinal} → (ab : odef A b) → b o≤ z
378 → IsSUP A (UnionCF A f mf ay supf b) ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf b) f b ) → supf b ≡ b 412 → IsSUP A (UnionCF A f mf ay supf b) ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf b) f b ) → supf b ≡ b
379 413
380 asupf : {x : Ordinal } → odef A (supf x) 414 asupf : {x : Ordinal } → odef A (supf x)
381 supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y 415 supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y
382 supf-< : {x y : Ordinal } → supf x o< supf y → supf x << supf y 416 supf-< : {x y : Ordinal } → supf x o< supf y → supf x << supf y
383 supfmax : {x : Ordinal } → z o< x → supf x ≡ supf z 417 supfmax : {x : Ordinal } → z o< x → supf x ≡ supf z
384 chain∋init : odef (UnionCF A f mf ay supf z) y 418
385 419 minsup : {x : Ordinal } → x o≤ z → MinSUP A (UnionCF A f mf ay supf x)
386 minsup : {x : Ordinal } → x o≤ z → MinSUP A (UnionCF A f mf ay supf x)
387 supf-is-minsup : {x : Ordinal } → (x≤z : x o≤ z) → supf x ≡ MinSUP.sup ( minsup x≤z ) 420 supf-is-minsup : {x : Ordinal } → (x≤z : x o≤ z) → supf x ≡ MinSUP.sup ( minsup x≤z )
388 csupf : {b : Ordinal } → supf b o< z → odef (UnionCF A f mf ay supf z) (supf b) -- supf z is not an element of this chain 421 csupf : {b : Ordinal } → supf b o< z → odef (UnionCF A f mf ay supf z) (supf b) -- supf z is not an element of this chain
389 422
390 chain : HOD 423 chain : HOD
391 chain = UnionCF A f mf ay supf z 424 chain = UnionCF A f mf ay supf z
392 chain⊆A : chain ⊆' A 425 chain⊆A : chain ⊆' A
393 chain⊆A = λ lt → proj1 lt 426 chain⊆A = λ lt → proj1 lt
394 427
395 sup : {x : Ordinal } → x o≤ z → SUP A (UnionCF A f mf ay supf x) 428 sup : {x : Ordinal } → x o≤ z → SUP A (UnionCF A f mf ay supf x)
396 sup {x} x≤z = M→S supf (minsup x≤z) 429 sup {x} x≤z = M→S supf (minsup x≤z)
397 430
398 s=ms : {x : Ordinal } → (x≤z : x o≤ z ) → & (SUP.sup (sup x≤z)) ≡ MinSUP.sup (minsup x≤z) 431 s=ms : {x : Ordinal } → (x≤z : x o≤ z ) → & (SUP.sup (sup x≤z)) ≡ MinSUP.sup (minsup x≤z)
399 s=ms {x} x≤z = &iso 432 s=ms {x} x≤z = &iso
400 433
434 chain∋init : odef chain y
435 chain∋init = ⟪ ay , ch-init (init ay refl) ⟫
401 f-next : {a z : Ordinal} → odef (UnionCF A f mf ay supf z) a → odef (UnionCF A f mf ay supf z) (f a) 436 f-next : {a z : Ordinal} → odef (UnionCF A f mf ay supf z) a → odef (UnionCF A f mf ay supf z) (f a)
437 f-next {a} ⟪ aa , ch-init fc ⟫ = ⟪ proj2 (mf a aa) , ch-init (fsuc _ fc) ⟫
402 f-next {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 ) ⟫ 438 f-next {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 ) ⟫
403 initial : {z : Ordinal } → odef chain z → * y ≤ * z 439 initial : {z : Ordinal } → odef chain z → * y ≤ * z
404 initial {a} ⟪ aa , ua ⟫ with ua 440 initial {a} ⟪ aa , ua ⟫ with ua
441 ... | ch-init fc = s≤fc y f mf fc
405 ... | ch-is-sup u u<x is-sup fc = ≤-ftrans (<=to≤ zc7) (s≤fc _ f mf fc) where 442 ... | ch-is-sup u u<x is-sup fc = ≤-ftrans (<=to≤ zc7) (s≤fc _ f mf fc) where
406 zc7 : y <= supf u 443 zc7 : y <= supf u
407 zc7 = ChainP.fcy<sup is-sup (init ay refl) 444 zc7 = ChainP.fcy<sup is-sup (init ay refl)
408 f-total : IsTotalOrderSet chain 445 f-total : IsTotalOrderSet chain
409 f-total {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where 446 f-total {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
410 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) 447 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
411 uz01 = chain-total A f mf ay supf ( (proj2 ca)) ( (proj2 cb)) 448 uz01 = chain-total A f mf ay supf ( (proj2 ca)) ( (proj2 cb))
412 449
413 supf-<= : {x y : Ordinal } → supf x <= supf y → supf x o≤ supf y 450 supf-<= : {x y : Ordinal } → supf x <= supf y → supf x o≤ supf y
414 supf-<= {x} {y} (case1 sx=sy) = o≤-refl0 sx=sy 451 supf-<= {x} {y} (case1 sx=sy) = o≤-refl0 sx=sy
415 supf-<= {x} {y} (case2 sx<sy) with trio< (supf x) (supf y) 452 supf-<= {x} {y} (case2 sx<sy) with trio< (supf x) (supf y)
416 ... | tri< a ¬b ¬c = o<→≤ a 453 ... | tri< a ¬b ¬c = o<→≤ a
417 ... | tri≈ ¬a b ¬c = o≤-refl0 b 454 ... | tri≈ ¬a b ¬c = o≤-refl0 b
418 ... | tri> ¬a ¬b c = ⊥-elim (<-irr (case2 sx<sy ) (supf-< c) ) 455 ... | tri> ¬a ¬b c = ⊥-elim (<-irr (case2 sx<sy ) (supf-< c) )
419 456
420 supf-inject : {x y : Ordinal } → supf x o< supf y → x o< y 457 supf-inject : {x y : Ordinal } → supf x o< supf y → x o< y
421 supf-inject {x} {y} sx<sy with trio< x y 458 supf-inject {x} {y} sx<sy with trio< x y
422 ... | tri< a ¬b ¬c = a 459 ... | tri< a ¬b ¬c = a
423 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy ) 460 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy )
424 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) ) 461 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) )
425 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy ) 462 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy )
426 ... | case2 lt = ⊥-elim ( o<> sx<sy lt ) 463 ... | case2 lt = ⊥-elim ( o<> sx<sy lt )
427 464
428 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 465 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
429 fcy<sup {u} {w} u≤z fc with chain∋init 466 fcy<sup {u} {w} u≤z fc with MinSUP.x≤sup (minsup u≤z) ⟪ subst (λ k → odef A k ) (sym &iso) (A∋fc {A} y f mf fc)
430 ... | ⟪ ay1 , ch-is-sup uy uy<x is-sup fcy ⟫ = <=-trans (ChainP.fcy<sup is-sup fc) ? 467 , ch-init (subst (λ k → FClosure A f y k) (sym &iso) fc ) ⟫
431 468 ... | case1 eq = case1 (subst (λ k → k ≡ supf u ) &iso (trans eq (sym (supf-is-minsup u≤z ) ) ))
432 -- with MinSUP.x≤sup (minsup u≤z) ⟪ subst (λ k → odef A k ) (sym &iso) (A∋fc {A} y f mf fc) 469 ... | case2 lt = case2 (subst₂ (λ j k → j << k ) &iso (sym (supf-is-minsup u≤z )) lt )
433 -- , ? ⟫ --ch-init (subst (λ k → FClosure A f y k) (sym &iso) fc ) ⟫ 470
434 -- ... | case1 eq = case1 (subst (λ k → k ≡ supf u ) &iso (trans eq (sym (supf-is-minsup u≤z ) ) )) 471 -- ordering is not proved here but in ZChain1
435 -- ... | case2 lt = case2 (subst₂ (λ j k → j << k ) &iso (sym (supf-is-minsup u≤z )) lt ) 472
436 473 IsMinSUP→NotHasPrev : {x sp : Ordinal } → odef A sp
437 -- ordering is not proved here but in ZChain1
438
439 IsMinSUP→NotHasPrev : {x sp : Ordinal } → odef A sp
440 → ({y : Ordinal} → odef (UnionCF A f mf ay supf x) y → (y ≡ sp ) ∨ (y << sp )) 474 → ({y : Ordinal} → odef (UnionCF A f mf ay supf x) y → (y ≡ sp ) ∨ (y << sp ))
441 → ( {a : Ordinal } → a << f a ) 475 → ( {a : Ordinal } → a << f a )
442 → ¬ ( HasPrev A (UnionCF A f mf ay supf x) f sp ) 476 → ¬ ( HasPrev A (UnionCF A f mf ay supf x) f sp )
443 IsMinSUP→NotHasPrev {x} {sp} asp is-sup <-mono-f hp = ⊥-elim (<-irr ( <=to≤ fsp≤sp) sp<fsp ) where 477 IsMinSUP→NotHasPrev {x} {sp} asp is-sup <-mono-f hp = ⊥-elim (<-irr ( <=to≤ fsp≤sp) sp<fsp ) where
444 sp<fsp : sp << f sp 478 sp<fsp : sp << f sp
445 sp<fsp = <-mono-f 479 sp<fsp = <-mono-f
446 pr = HasPrev.y hp 480 pr = HasPrev.y hp
447 im00 : f (f pr) <= sp 481 im00 : f (f pr) <= sp
448 im00 = is-sup ( f-next (f-next (HasPrev.ay hp))) 482 im00 = is-sup ( f-next (f-next (HasPrev.ay hp)))
449 fsp≤sp : f sp <= sp 483 fsp≤sp : f sp <= sp
450 fsp≤sp = subst (λ k → f k <= sp ) (sym (HasPrev.x=fy hp)) im00 484 fsp≤sp = subst (λ k → f k <= sp ) (sym (HasPrev.x=fy hp)) im00
451 485
452 record ZChain1 ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) 486 UChain⊆ : { supf1 : Ordinal → Ordinal }
487 → ( { x : Ordinal } → x o< z → supf x ≡ supf1 x)
488 → ( { x : Ordinal } → z o≤ x → supf z o≤ supf1 x)
489 → UnionCF A f mf ay supf z ⊆' UnionCF A f mf ay supf1 z
490 UChain⊆ {supf1} eq<x lex ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
491 UChain⊆ {supf1} eq<x lex ⟪ az , ch-is-sup u {x} u<x is-sup fc ⟫ = ⟪ az , ch-is-sup u u<x1 cp1 fc1 ⟫ where
492 u<x0 : u o< z
493 u<x0 = supf-inject u<x
494 u<x1 : supf1 u o< supf1 z
495 u<x1 = subst (λ k → k o< supf1 z ) (eq<x u<x0) (ordtrans<-≤ u<x (lex o≤-refl ) )
496 fc1 : FClosure A f (supf1 u) x
497 fc1 = subst (λ k → FClosure A f k x ) (eq<x u<x0) fc
498 uc01 : {s : Ordinal } → supf1 s o< supf1 u → s o< z
499 uc01 {s} s<u with trio< s z
500 ... | tri< a ¬b ¬c = a
501 ... | tri≈ ¬a b ¬c = ⊥-elim ( o≤> uc02 s<u ) where -- (supf-mono (o<→≤ u<x0))
502 uc02 : supf1 u o≤ supf1 s
503 uc02 = begin
504 supf1 u <⟨ u<x1 ⟩
505 supf1 z ≡⟨ cong supf1 (sym b) ⟩
506 supf1 s ∎ where open o≤-Reasoning O
507 ... | tri> ¬a ¬b c = ⊥-elim ( o≤> uc03 s<u ) where
508 uc03 : supf1 u o≤ supf1 s
509 uc03 = begin
510 supf1 u ≡⟨ sym (eq<x u<x0) ⟩
511 supf u <⟨ u<x ⟩
512 supf z ≤⟨ lex (o<→≤ c) ⟩
513 supf1 s ∎ where open o≤-Reasoning O
514 cp1 : ChainP A f mf ay supf1 u
515 cp1 = record { fcy<sup = λ {z} fc → subst (λ k → (z ≡ k) ∨ ( z << k ) ) (eq<x u<x0) (ChainP.fcy<sup is-sup fc )
516 ; order = λ {s} {z} s<u fc → subst (λ k → (z ≡ k) ∨ ( z << k ) ) (eq<x u<x0)
517 (ChainP.order is-sup (subst₂ (λ j k → j o< k ) (sym (eq<x (uc01 s<u) )) (sym (eq<x u<x0)) s<u)
518 (subst (λ k → FClosure A f k z ) (sym (eq<x (uc01 s<u) )) fc ))
519 ; supu=u = trans (sym (eq<x u<x0)) (ChainP.supu=u is-sup) }
520
521 record ZChain1 ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)
453 {y : Ordinal} (ay : odef A y) (zc : ZChain A f mf ay (& A)) ( z : Ordinal ) : Set (Level.suc n) where 522 {y : Ordinal} (ay : odef A y) (zc : ZChain A f mf ay (& A)) ( z : Ordinal ) : Set (Level.suc n) where
454 supf = ZChain.supf zc 523 supf = ZChain.supf zc
455 field 524 field
456 is-max : {a b : Ordinal } → (ca : odef (UnionCF A f mf ay supf z) a ) → supf b o< supf z → (ab : odef A b) 525 is-max : {a b : Ordinal } → (ca : odef (UnionCF A f mf ay supf z) a ) → supf b o< supf z → (ab : odef A b)
457 → HasPrev A (UnionCF A f mf ay supf z) f b ∨ IsSUP A (UnionCF A f mf ay supf b) ab 526 → HasPrev A (UnionCF A f mf ay supf z) f b ∨ IsSUP A (UnionCF A f mf ay supf b) ab
458 → * a < * b → odef ((UnionCF A f mf ay supf z)) b 527 → * a < * b → odef ((UnionCF A f mf ay supf z)) b
459 order : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → (z1 ≡ supf b) ∨ (z1 << supf b) 528 order : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → (z1 ≡ supf b) ∨ (z1 << supf b)
460 529
461 record Maximal ( A : HOD ) : Set (Level.suc n) where 530 record Maximal ( A : HOD ) : Set (Level.suc n) where
462 field 531 field
463 maximal : HOD 532 maximal : HOD
464 as : A ∋ maximal 533 as : A ∋ maximal
465 ¬maximal<x : {x : HOD} → A ∋ x → ¬ maximal < x -- A is Partial, use negative 534 ¬maximal<x : {x : HOD} → A ∋ x → ¬ maximal < x -- A is Partial, use negative
466 535
467 Zorn-lemma : { A : HOD } 536 init-uchain : (A : HOD) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } → (ay : odef A y )
468 → o∅ o< & A 537 { supf : Ordinal → Ordinal } { x : Ordinal } → odef (UnionCF A f mf ay supf x) y
538 init-uchain A f mf ay = ⟪ ay , ch-init (init ay refl) ⟫
539
540 Zorn-lemma : { A : HOD }
541 → o∅ o< & A
469 → ( ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → SUP A B ) -- SUP condition 542 → ( ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
470 → Maximal A 543 → Maximal A
471 Zorn-lemma {A} 0<A supP = zorn00 where 544 Zorn-lemma {A} 0<A supP = zorn00 where
472 <-irr0 : {a b : HOD} → A ∋ a → A ∋ b → (a ≡ b ) ∨ (a < b ) → b < a → ⊥ 545 <-irr0 : {a b : HOD} → A ∋ a → A ∋ b → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
473 <-irr0 {a} {b} A∋a A∋b = <-irr 546 <-irr0 {a} {b} A∋a A∋b = <-irr
474 z07 : {y : Ordinal} {A : HOD } → {P : Set n} → odef A y ∧ P → y o< & A 547 z07 : {y : Ordinal} {A : HOD } → {P : Set n} → odef A y ∧ P → y o< & A
475 z07 {y} {A} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p ))) 548 z07 {y} {A} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
476 s : HOD 549 s : HOD
477 s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) 550 s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A ))
478 as : A ∋ * ( & s ) 551 as : A ∋ * ( & s )
479 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 )) ) 552 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 )) )
480 as0 : odef A (& s ) 553 as0 : odef A (& s )
481 as0 = subst (λ k → odef A k ) &iso as 554 as0 = subst (λ k → odef A k ) &iso as
482 s<A : & s o< & A 555 s<A : & s o< & A
483 s<A = c<→o< (subst (λ k → odef A (& k) ) *iso as ) 556 s<A = c<→o< (subst (λ k → odef A (& k) ) *iso as )
484 HasMaximal : HOD 557 HasMaximal : HOD
485 HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) → odef A m → ¬ (* x < * m)) } ; odmax = & A ; <odmax = z07 } 558 HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) → odef A m → ¬ (* x < * m)) } ; odmax = & A ; <odmax = z07 }
486 no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → odef A x ∧ ((m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m) )) → ⊥ 559 no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → odef A x ∧ ((m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m) )) → ⊥
487 no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ ) 560 no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ )
488 Gtx : { x : HOD} → A ∋ x → HOD 561 Gtx : { x : HOD} → A ∋ x → HOD
489 Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 } 562 Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 }
490 z08 : ¬ Maximal A → HasMaximal =h= od∅ 563 z08 : ¬ Maximal A → HasMaximal =h= od∅
491 z08 nmx = record { eq→ = λ {x} lt → ⊥-elim ( nmx record {maximal = * x ; as = subst (λ k → odef A k) (sym &iso) (proj1 lt) 564 z08 nmx = record { eq→ = λ {x} lt → ⊥-elim ( nmx record {maximal = * x ; as = subst (λ k → odef A k) (sym &iso) (proj1 lt)
492 ; ¬maximal<x = λ {y} ay → subst (λ k → ¬ (* x < k)) *iso (proj2 lt (& y) ay) } ) ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )} 565 ; ¬maximal<x = λ {y} ay → subst (λ k → ¬ (* x < k)) *iso (proj2 lt (& y) ay) } ) ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )}
493 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)) 566 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))
494 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 567 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
495 ¬x<m : ¬ (* x < * m) 568 ¬x<m : ¬ (* x < * m)
496 ¬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) 569 ¬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)
497 570
498 minsupP : ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → MinSUP A B 571 minsupP : ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → MinSUP A B
499 minsupP B B⊆A total = m02 where 572 minsupP B B⊆A total = m02 where
500 xsup : (sup : Ordinal ) → Set n 573 xsup : (sup : Ordinal ) → Set n
501 xsup sup = {w : Ordinal } → odef B w → (w ≡ sup ) ∨ (w << sup ) 574 xsup sup = {w : Ordinal } → odef B w → (w ≡ sup ) ∨ (w << sup )
502 ∀-imply-or : {A : Ordinal → Set n } {B : Set n } 575 ∀-imply-or : {A : Ordinal → Set n } {B : Set n }
503 → ((x : Ordinal ) → A x ∨ B) → ((x : Ordinal ) → A x) ∨ B 576 → ((x : Ordinal ) → A x ∨ B) → ((x : Ordinal ) → A x) ∨ B
534 s1 = & (SUP.sup S) 607 s1 = & (SUP.sup S)
535 m05 : {w : Ordinal } → odef B w → (w ≡ s1 ) ∨ (w << s1 ) 608 m05 : {w : Ordinal } → odef B w → (w ≡ s1 ) ∨ (w << s1 )
536 m05 {w} bw with SUP.x≤sup S {* w} (subst (λ k → odef B k) (sym &iso) bw ) 609 m05 {w} bw with SUP.x≤sup S {* w} (subst (λ k → odef B k) (sym &iso) bw )
537 ... | case1 eq = case1 ( subst₂ (λ j k → j ≡ k ) &iso refl (cong (&) eq) ) 610 ... | case1 eq = case1 ( subst₂ (λ j k → j ≡ k ) &iso refl (cong (&) eq) )
538 ... | case2 lt = case2 ( subst (λ k → _ < k ) (sym *iso) lt ) 611 ... | case2 lt = case2 ( subst (λ k → _ < k ) (sym *iso) lt )
539 m02 : MinSUP A B 612 m02 : MinSUP A B
540 m02 = dont-or (m00 (& A)) m03 613 m02 = dont-or (m00 (& A)) m03
541 614
542 -- Uncountable ascending chain by axiom of choice 615 -- Uncountable ascending chain by axiom of choice
543 cf : ¬ Maximal A → Ordinal → Ordinal 616 cf : ¬ Maximal A → Ordinal → Ordinal
544 cf nmx x with ODC.∋-p O A (* x) 617 cf nmx x with ODC.∋-p O A (* x)
566 -- maximality of chain 639 -- maximality of chain
567 -- 640 --
568 -- supf is fixed for z ≡ & A , we can prove order and is-max 641 -- supf is fixed for z ≡ & A , we can prove order and is-max
569 -- 642 --
570 643
571 SZ1 : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) 644 SZ1 : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)
572 {init : Ordinal} (ay : odef A init) (zc : ZChain A f mf ay (& A)) (x : Ordinal) → ZChain1 A f mf ay zc x 645 {init : Ordinal} (ay : odef A init) (zc : ZChain A f mf ay (& A)) (x : Ordinal) → ZChain1 A f mf ay zc x
573 SZ1 f mf {y} ay zc x = zc1 x where 646 SZ1 f mf {y} ay zc x = zc1 x where
574 chain-mono1 : {a b c : Ordinal} → a o≤ b 647 chain-mono1 : {a b c : Ordinal} → a o≤ b
575 → odef (UnionCF A f mf ay (ZChain.supf zc) a) c → odef (UnionCF A f mf ay (ZChain.supf zc) b) c 648 → odef (UnionCF A f mf ay (ZChain.supf zc) a) c → odef (UnionCF A f mf ay (ZChain.supf zc) b) c
576 chain-mono1 {a} {b} {c} a≤b = chain-mono f mf ay (ZChain.supf zc) (ZChain.supf-mono zc) a≤b 649 chain-mono1 {a} {b} {c} a≤b = chain-mono f mf ay (ZChain.supf zc) (ZChain.supf-mono zc) a≤b
577 is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a → (ab : odef A b) 650 is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a → (ab : odef A b)
578 → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) f b 651 → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) f b
579 → * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b 652 → * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b
580 is-max-hp x {a} {b} ua ab has-prev a<b with HasPrev.ay has-prev 653 is-max-hp x {a} {b} ua ab has-prev a<b with HasPrev.ay has-prev
654 ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫
581 ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ⟪ ab , subst (λ k → UChain A f mf ay (ZChain.supf zc) x k ) 655 ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ⟪ ab , subst (λ k → UChain A f mf ay (ZChain.supf zc) x k )
582 (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc)) ⟫ 656 (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc)) ⟫
583 657
584 supf = ZChain.supf zc 658 supf = ZChain.supf zc
585 659
586 csupf-fc : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → UnionCF A f mf ay supf b ∋ * z1 660 csupf-fc : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → UnionCF A f mf ay supf b ∋ * z1
587 csupf-fc {b} {s} {z1} b<z ss<sb (init x refl ) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc05 where 661 csupf-fc {b} {s} {z1} b<z ss<sb (init x refl ) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc05 where
591 s<z = ordtrans s<b b<z 665 s<z = ordtrans s<b b<z
592 zc04 : odef (UnionCF A f mf ay supf (& A)) (supf s) 666 zc04 : odef (UnionCF A f mf ay supf (& A)) (supf s)
593 zc04 = ZChain.csupf zc (z09 (ZChain.asupf zc)) 667 zc04 = ZChain.csupf zc (z09 (ZChain.asupf zc))
594 zc05 : odef (UnionCF A f mf ay supf b) (supf s) 668 zc05 : odef (UnionCF A f mf ay supf b) (supf s)
595 zc05 with zc04 669 zc05 with zc04
670 ... | ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫
596 ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ as , ch-is-sup u zc08 is-sup fc ⟫ where 671 ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ as , ch-is-sup u zc08 is-sup fc ⟫ where
597 zc07 : FClosure A f (supf u) (supf s) -- supf u ≤ supf s → supf u o≤ supf s 672 zc07 : FClosure A f (supf u) (supf s) -- supf u ≤ supf s → supf u o≤ supf s
598 zc07 = fc 673 zc07 = fc
599 zc06 : supf u ≡ u 674 zc06 : supf u ≡ u
600 zc06 = ChainP.supu=u is-sup 675 zc06 = ChainP.supu=u is-sup
601 zc08 : supf u o< supf b 676 zc08 : supf u o< supf b
602 zc08 = ordtrans≤-< (ZChain.supf-<= zc (≤to<= ( s≤fc _ f mf fc ))) ss<sb 677 zc08 = ordtrans≤-< (ZChain.supf-<= zc (≤to<= ( s≤fc _ f mf fc ))) ss<sb
603 csupf-fc {b} {s} {z1} b<z ss≤sb (fsuc x fc) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc04 where 678 csupf-fc {b} {s} {z1} b<z ss≤sb (fsuc x fc) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc04 where
604 zc04 : odef (UnionCF A f mf ay supf b) (f x) 679 zc04 : odef (UnionCF A f mf ay supf b) (f x)
605 zc04 with subst (λ k → odef (UnionCF A f mf ay supf b) k ) &iso (csupf-fc b<z ss≤sb fc ) 680 zc04 with subst (λ k → odef (UnionCF A f mf ay supf b) k ) &iso (csupf-fc b<z ss≤sb fc )
606 ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ proj2 (mf _ as) , ch-is-sup u u<x is-sup (fsuc _ fc) ⟫ 681 ... | ⟪ as , ch-init fc ⟫ = ⟪ proj2 (mf _ as) , ch-init (fsuc _ fc) ⟫
682 ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ proj2 (mf _ as) , ch-is-sup u u<x is-sup (fsuc _ fc) ⟫
607 order : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → (z1 ≡ supf b) ∨ (z1 << supf b) 683 order : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → (z1 ≡ supf b) ∨ (z1 << supf b)
608 order {b} {s} {z1} b<z ss<sb fc = zc04 where 684 order {b} {s} {z1} b<z ss<sb fc = zc04 where
609 zc00 : ( z1 ≡ MinSUP.sup (ZChain.minsup zc (o<→≤ b<z) )) ∨ ( z1 << MinSUP.sup ( ZChain.minsup zc (o<→≤ b<z) ) ) 685 zc00 : ( z1 ≡ MinSUP.sup (ZChain.minsup zc (o<→≤ b<z) )) ∨ ( z1 << MinSUP.sup ( ZChain.minsup zc (o<→≤ b<z) ) )
610 zc00 = MinSUP.x≤sup (ZChain.minsup zc (o<→≤ b<z) ) (subst (λ k → odef (UnionCF A f mf ay (ZChain.supf zc) b) k ) &iso (csupf-fc b<z ss<sb fc )) 686 zc00 = MinSUP.x≤sup (ZChain.minsup zc (o<→≤ b<z) ) (subst (λ k → odef (UnionCF A f mf ay (ZChain.supf zc) b) k ) &iso (csupf-fc b<z ss<sb fc ))
611 -- supf (supf b) ≡ supf b 687 -- supf (supf b) ≡ supf b
623 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a → 699 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a →
624 ZChain.supf zc b o< ZChain.supf zc x → (ab : odef A b) → 700 ZChain.supf zc b o< ZChain.supf zc x → (ab : odef A b) →
625 HasPrev A (UnionCF A f mf ay supf x) f b ∨ IsSUP A (UnionCF A f mf ay supf b) ab → 701 HasPrev A (UnionCF A f mf ay supf x) f b ∨ IsSUP A (UnionCF A f mf ay supf b) ab →
626 * a < * b → odef (UnionCF A f mf ay supf x) b 702 * a < * b → odef (UnionCF A f mf ay supf x) b
627 is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P 703 is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P
628 is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b 704 is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
629 is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup 705 is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup
630 = ⟪ ab , ch-is-sup b sb<sx m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫ where 706 = ⟪ ab , ch-is-sup b sb<sx m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫ where
631 b<A : b o< & A 707 b<A : b o< & A
632 b<A = z09 ab 708 b<A = z09 ab
633 b<x : b o< x 709 b<x : b o< x
634 b<x = ZChain.supf-inject zc sb<sx 710 b<x = ZChain.supf-inject zc sb<sx
635 m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) f b 711 m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) f b
636 m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay = 712 m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay =
637 chain-mono1 (o<→≤ b<x) (HasPrev.ay nhp) ; x=fy = HasPrev.x=fy nhp } ) 713 chain-mono1 (o<→≤ b<x) (HasPrev.ay nhp) ; x=fy = HasPrev.x=fy nhp } )
638 m05 : ZChain.supf zc b ≡ b 714 m05 : ZChain.supf zc b ≡ b
639 m05 = ZChain.sup=u zc ab (o<→≤ (z09 ab) ) ⟪ record { x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz } , m04 ⟫ 715 m05 = ZChain.sup=u zc ab (o<→≤ (z09 ab) ) ⟪ record { x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz } , m04 ⟫
640 m08 : {z : Ordinal} → (fcz : FClosure A f y z ) → z <= ZChain.supf zc b 716 m08 : {z : Ordinal} → (fcz : FClosure A f y z ) → z <= ZChain.supf zc b
641 m08 {z} fcz = ZChain.fcy<sup zc (o<→≤ b<A) fcz 717 m08 {z} fcz = ZChain.fcy<sup zc (o<→≤ b<A) fcz
642 m09 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b 718 m09 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b
643 → FClosure A f (ZChain.supf zc s) z1 → z1 <= ZChain.supf zc b 719 → FClosure A f (ZChain.supf zc s) z1 → z1 <= ZChain.supf zc b
644 m09 {s} {z} s<b fcz = order b<A s<b fcz 720 m09 {s} {z} s<b fcz = order b<A s<b fcz
645 m06 : ChainP A f mf ay supf b 721 m06 : ChainP A f mf ay supf b
646 m06 = record { fcy<sup = m08 ; order = m09 ; supu=u = m05 } 722 m06 = record { fcy<sup = m08 ; order = m09 ; supu=u = m05 }
647 ... | no lim = record { is-max = is-max ; order = order } where 723 ... | no lim = record { is-max = is-max ; order = order } where
648 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a → 724 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a →
649 ZChain.supf zc b o< ZChain.supf zc x → (ab : odef A b) → 725 ZChain.supf zc b o< ZChain.supf zc x → (ab : odef A b) →
650 HasPrev A (UnionCF A f mf ay supf x) f b ∨ IsSUP A (UnionCF A f mf ay supf b) ab → 726 HasPrev A (UnionCF A f mf ay supf x) f b ∨ IsSUP A (UnionCF A f mf ay supf b) ab →
651 * a < * b → odef (UnionCF A f mf ay supf x) b 727 * a < * b → odef (UnionCF A f mf ay supf x) b
652 is-max {a} {b} ua sb<sx ab P a<b with ODC.or-exclude O P 728 is-max {a} {b} ua sb<sx ab P a<b with ODC.or-exclude O P
653 is-max {a} {b} ua sb<sx ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b 729 is-max {a} {b} ua sb<sx ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
654 is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup 730 is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup with IsSUP.x≤sup (proj2 is-sup) (init-uchain A f mf ay )
655 = ⟪ ab , ch-is-sup b sb<sx m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫ where 731 ... | 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 )) ⟫
732 ... | case2 y<b = ⟪ ab , ch-is-sup b sb<sx m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫ where
656 m09 : b o< & A 733 m09 : b o< & A
657 m09 = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab)) 734 m09 = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
658 b<x : b o< x 735 b<x : b o< x
659 b<x = ZChain.supf-inject zc sb<sx 736 b<x = ZChain.supf-inject zc sb<sx
660 m07 : {z : Ordinal} → FClosure A f y z → z <= ZChain.supf zc b 737 m07 : {z : Ordinal} → FClosure A f y z → z <= ZChain.supf zc b
661 m07 {z} fc = ZChain.fcy<sup zc (o<→≤ m09) fc 738 m07 {z} fc = ZChain.fcy<sup zc (o<→≤ m09) fc
662 m08 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b 739 m08 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b
663 → FClosure A f (ZChain.supf zc s) z1 → z1 <= ZChain.supf zc b 740 → FClosure A f (ZChain.supf zc s) z1 → z1 <= ZChain.supf zc b
664 m08 {s} {z1} s<b fc = order m09 s<b fc 741 m08 {s} {z1} s<b fc = order m09 s<b fc
665 m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) f b 742 m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) f b
666 m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay = 743 m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay =
667 chain-mono1 (o<→≤ b<x) (HasPrev.ay nhp) 744 chain-mono1 (o<→≤ b<x) (HasPrev.ay nhp)
668 ; x=fy = HasPrev.x=fy nhp } ) 745 ; x=fy = HasPrev.x=fy nhp } )
669 m05 : ZChain.supf zc b ≡ b 746 m05 : ZChain.supf zc b ≡ b
670 m05 = ZChain.sup=u zc ab (o<→≤ m09) ⟪ record { x≤sup = λ lt → IsSUP.x≤sup (proj2 is-sup) lt } , m04 ⟫ -- ZChain on x 747 m05 = ZChain.sup=u zc ab (o<→≤ m09) ⟪ record { x≤sup = λ lt → IsSUP.x≤sup (proj2 is-sup) lt } , m04 ⟫ -- ZChain on x
671 m06 : ChainP A f mf ay supf b 748 m06 : ChainP A f mf ay supf b
672 m06 = record { fcy<sup = m07 ; order = m08 ; supu=u = m05 } 749 m06 = record { fcy<sup = m07 ; order = m08 ; supu=u = m05 }
673 750
674 uchain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → HOD 751 uchain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → HOD
675 uchain f mf {y} ay = record { od = record { def = λ x → FClosure A f y x } ; odmax = & A ; <odmax = 752 uchain f mf {y} ay = record { od = record { def = λ x → FClosure A f y x } ; odmax = & A ; <odmax =
676 λ {z} cz → subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (A∋fc y f mf cz ))) } 753 λ {z} cz → subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (A∋fc y f mf cz ))) }
677 754
678 utotal : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) 755 utotal : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
679 → IsTotalOrderSet (uchain f mf ay) 756 → IsTotalOrderSet (uchain f mf ay)
680 utotal f mf {y} ay {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where 757 utotal f mf {y} ay {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
681 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) 758 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
682 uz01 = fcn-cmp y f mf ca cb 759 uz01 = fcn-cmp y f mf ca cb
683 760
684 ysup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) 761 ysup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
685 → MinSUP A (uchain f mf ay) 762 → MinSUP A (uchain f mf ay)
686 ysup f mf {y} ay = minsupP (uchain f mf ay) (λ lt → A∋fc y f mf lt) (utotal f mf ay) 763 ysup f mf {y} ay = minsupP (uchain f mf ay) (λ lt → A∋fc y f mf lt) (utotal f mf ay)
687 764
688 UChainPreserve : {x z : Ordinal } { f : Ordinal → Ordinal } → {mf : ≤-monotonic-f A f } {y : Ordinal} {ay : odef A y}
689 → { supf0 supf1 : Ordinal → Ordinal }
690 → ( ( z : Ordinal ) → z o< x → supf0 z ≡ supf1 z)
691 → UnionCF A f mf ay supf0 x ≡ UnionCF A f mf ay supf1 x
692 UChainPreserve {x} {f} {mf} {y} {ay} {supf0} {supf1} <x→eq = ?
693 765
694 SUP⊆ : { B C : HOD } → B ⊆' C → SUP A C → SUP A B 766 SUP⊆ : { B C : HOD } → B ⊆' C → SUP A C → SUP A B
695 SUP⊆ {B} {C} B⊆C sup = record { sup = SUP.sup sup ; as = SUP.as sup ; x≤sup = λ lt → SUP.x≤sup sup (B⊆C lt) } 767 SUP⊆ {B} {C} B⊆C sup = record { sup = SUP.sup sup ; as = SUP.as sup ; x≤sup = λ lt → SUP.x≤sup sup (B⊆C lt) }
696 768
697 record xSUP (B : HOD) (f : Ordinal → Ordinal ) (x : Ordinal) : Set n where 769 record xSUP (B : HOD) (f : Ordinal → Ordinal ) (x : Ordinal) : Set n where
701 773
702 -- 774 --
703 -- create all ZChains under o< x 775 -- create all ZChains under o< x
704 -- 776 --
705 777
706 ind : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal) 778 ind : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal)
707 → ((z : Ordinal) → z o< x → ZChain A f mf ay z) → ZChain A f mf ay x 779 → ((z : Ordinal) → z o< x → ZChain A f mf ay z) → ZChain A f mf ay x
708 ind f mf {y} ay x prev with Oprev-p x 780 ind f mf {y} ay x prev with Oprev-p x
709 ... | yes op = zc41 where 781 ... | yes op = zc41 where
710 -- 782 --
711 -- we have previous ordinal to use induction 783 -- we have previous ordinal to use induction
712 -- 784 --
713 px = Oprev.oprev op 785 px = Oprev.oprev op
714 zc : ZChain A f mf ay (Oprev.oprev op) 786 zc : ZChain A f mf ay (Oprev.oprev op)
715 zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) 787 zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc )
716 px<x : px o< x 788 px<x : px o< x
717 px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc 789 px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc
718 opx=x : osuc px ≡ x 790 opx=x : osuc px ≡ x
719 opx=x = Oprev.oprev=x op 791 opx=x = Oprev.oprev=x op
720 792
721 zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px 793 zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px
722 zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt 794 zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt
723 795
724 supf0 = ZChain.supf zc 796 supf0 = ZChain.supf zc
725 pchain : HOD 797 pchain : HOD
726 pchain = UnionCF A f mf ay supf0 px 798 pchain = UnionCF A f mf ay supf0 px
727 799
728 supf-mono : {a b : Ordinal } → a o≤ b → supf0 a o≤ supf0 b 800 supf-mono : {a b : Ordinal } → a o≤ b → supf0 a o≤ supf0 b
729 supf-mono = ZChain.supf-mono zc 801 supf-mono = ZChain.supf-mono zc
730 802
731 zc04 : {b : Ordinal} → b o≤ x → (b o≤ px ) ∨ (b ≡ x ) 803 zc04 : {b : Ordinal} → b o≤ x → (b o≤ px ) ∨ (b ≡ x )
732 zc04 {b} b≤x with trio< b px 804 zc04 {b} b≤x with trio< b px
733 ... | tri< a ¬b ¬c = case1 (o<→≤ a) 805 ... | tri< a ¬b ¬c = case1 (o<→≤ a)
734 ... | tri≈ ¬a b ¬c = case1 (o≤-refl0 b) 806 ... | tri≈ ¬a b ¬c = case1 (o≤-refl0 b)
735 ... | tri> ¬a ¬b px<b with osuc-≡< b≤x 807 ... | tri> ¬a ¬b px<b with osuc-≡< b≤x
736 ... | case1 eq = case2 eq 808 ... | case1 eq = case2 eq
737 ... | case2 b<x = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ ) 809 ... | case2 b<x = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ )
738 810
739 -- 811 --
740 -- find the next value of supf 812 -- find the next value of supf
741 -- 813 --
742 814
743 pchainpx : HOD 815 pchainpx : HOD
744 pchainpx = record { od = record { def = λ z → (odef A z ∧ UChain A f mf ay supf0 px z ) 816 pchainpx = record { od = record { def = λ z → (odef A z ∧ UChain A f mf ay supf0 px z )
745 ∨ FClosure A f (supf0 px) z } ; odmax = & A ; <odmax = zc00 } where 817 ∨ FClosure A f (supf0 px) z } ; odmax = & A ; <odmax = zc00 } where
746 zc00 : {z : Ordinal } → (odef A z ∧ UChain A f mf ay supf0 px z ) ∨ FClosure A f (supf0 px) z → z o< & A 818 zc00 : {z : Ordinal } → (odef A z ∧ UChain A f mf ay supf0 px z ) ∨ FClosure A f (supf0 px) z → z o< & A
747 zc00 {z} (case1 lt) = z07 lt 819 zc00 {z} (case1 lt) = z07 lt
748 zc00 {z} (case2 fc) = z09 ( A∋fc (supf0 px) f mf fc ) 820 zc00 {z} (case2 fc) = z09 ( A∋fc (supf0 px) f mf fc )
749 821
750 zc02 : { a b : Ordinal } → odef A a ∧ UChain A f mf ay supf0 px a → FClosure A f (supf0 px) b → a <= b 822 zc02 : { a b : Ordinal } → odef A a ∧ UChain A f mf ay supf0 px a → FClosure A f (supf0 px) b → a <= b
751 zc02 {a} {b} ca fb = zc05 fb where 823 zc02 {a} {b} ca fb = zc05 fb where
752 zc06 : MinSUP.sup (ZChain.minsup zc o≤-refl) ≡ supf0 px 824 zc06 : MinSUP.sup (ZChain.minsup zc o≤-refl) ≡ supf0 px
753 zc06 = trans (sym ( ZChain.supf-is-minsup zc o≤-refl )) refl 825 zc06 = trans (sym ( ZChain.supf-is-minsup zc o≤-refl )) refl
754 zc05 : {b : Ordinal } → FClosure A f (supf0 px) b → a <= b 826 zc05 : {b : Ordinal } → FClosure A f (supf0 px) b → a <= b
755 zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc (supf0 px) f mf fb )) 827 zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc (supf0 px) f mf fb ))
756 ... | case1 eq = subst (λ k → a <= k ) (subst₂ (λ j k → j ≡ k) &iso &iso (cong (&) eq)) (zc05 fb) 828 ... | case1 eq = subst (λ k → a <= k ) (subst₂ (λ j k → j ≡ k) &iso &iso (cong (&) eq)) (zc05 fb)
757 ... | case2 lt = <=-trans (zc05 fb) (case2 lt) 829 ... | case2 lt = <=-trans (zc05 fb) (case2 lt)
758 zc05 (init b1 refl) with MinSUP.x≤sup (ZChain.minsup zc o≤-refl) 830 zc05 (init b1 refl) with MinSUP.x≤sup (ZChain.minsup zc o≤-refl)
759 (subst (λ k → odef A k ∧ UChain A f mf ay supf0 px k) (sym &iso) ca ) 831 (subst (λ k → odef A k ∧ UChain A f mf ay supf0 px k) (sym &iso) ca )
760 ... | case1 eq = case1 (subst₂ (λ j k → j ≡ k ) &iso zc06 eq ) 832 ... | case1 eq = case1 (subst₂ (λ j k → j ≡ k ) &iso zc06 eq )
761 ... | case2 lt = case2 (subst₂ (λ j k → j < k ) *iso (cong (*) zc06) lt ) 833 ... | case2 lt = case2 (subst₂ (λ j k → j < k ) *iso (cong (*) zc06) lt )
762 834
763 ptotal : IsTotalOrderSet pchainpx 835 ptotal : IsTotalOrderSet pchainpx
764 ptotal (case1 a) (case1 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso 836 ptotal (case1 a) (case1 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso
765 (chain-total A f mf ay supf0 (proj2 a) (proj2 b)) 837 (chain-total A f mf ay supf0 (proj2 a) (proj2 b))
766 ptotal {a0} {b0} (case1 a) (case2 b) with zc02 a b 838 ptotal {a0} {b0} (case1 a) (case2 b) with zc02 a b
767 ... | case1 eq = tri≈ (<-irr (case1 (sym eq1))) eq1 (<-irr (case1 eq1)) where 839 ... | case1 eq = tri≈ (<-irr (case1 (sym eq1))) eq1 (<-irr (case1 eq1)) where
768 eq1 : a0 ≡ b0 840 eq1 : a0 ≡ b0
769 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq ) 841 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
770 ... | case2 lt = tri< lt1 (λ eq → <-irr (case1 (sym eq)) lt1) (<-irr (case2 lt1)) where 842 ... | case2 lt = tri< lt1 (λ eq → <-irr (case1 (sym eq)) lt1) (<-irr (case2 lt1)) where
776 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq ) 848 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
777 ... | case2 lt = tri> (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where 849 ... | case2 lt = tri> (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where
778 lt1 : a0 < b0 850 lt1 : a0 < b0
779 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt 851 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt
780 ptotal (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp (supf0 px) f mf a b) 852 ptotal (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp (supf0 px) f mf a b)
781 853
782 pcha : pchainpx ⊆' A 854 pcha : pchainpx ⊆' A
783 pcha (case1 lt) = proj1 lt 855 pcha (case1 lt) = proj1 lt
784 pcha (case2 fc) = A∋fc _ f mf fc 856 pcha (case2 fc) = A∋fc _ f mf fc
785 857
786 sup1 : MinSUP A pchainpx 858 sup1 : MinSUP A pchainpx
787 sup1 = minsupP pchainpx pcha ptotal 859 sup1 = minsupP pchainpx pcha ptotal
788 sp1 = MinSUP.sup sup1 860 sp1 = MinSUP.sup sup1
789 861
790 -- 862 --
791 -- supf0 px o≤ sp1 863 -- supf0 px o≤ sp1
792 -- 864 --
793 865
794 zc41 : ZChain A f mf ay x 866 zc41 : ZChain A f mf ay x
795 zc41 with MinSUP.x≤sup sup1 (case2 (init (ZChain.asupf zc {px}) refl )) 867 zc41 with MinSUP.x≤sup sup1 (case2 (init (ZChain.asupf zc {px}) refl ))
796 zc41 | (case2 sfpx<sp1) = record { supf = supf1 ; sup=u = ? ; asupf = ? ; supf-mono = supf1-mono ; supf-< = ? 868 zc41 | (case2 sfpx<sp1) = record { supf = supf1 ; sup=u = ? ; asupf = ? ; supf-mono = supf1-mono ; supf-< = ?
797 ; supfmax = ? ; minsup = ? ; supf-is-minsup = ? ; csupf = csupf1 } where 869 ; supfmax = ? ; minsup = ? ; supf-is-minsup = ? ; csupf = csupf1 } where
798 -- supf0 px is included by the chain of sp1 870 -- supf0 px is included by the chain of sp1
799 -- ( UnionCF A f mf ay supf0 px ∪ FClosure (supf0 px) ) ≡ UnionCF supf1 x 871 -- ( UnionCF A f mf ay supf0 px ∪ FClosure (supf0 px) ) ≡ UnionCF supf1 x
800 -- supf1 x ≡ sp1, which is not included now 872 -- supf1 x ≡ sp1, which is not included now
801 873
802 supf1 : Ordinal → Ordinal 874 supf1 : Ordinal → Ordinal
803 supf1 z with trio< z px 875 supf1 z with trio< z px
804 ... | tri< a ¬b ¬c = supf0 z 876 ... | tri< a ¬b ¬c = supf0 z
805 ... | tri≈ ¬a b ¬c = supf0 z 877 ... | tri≈ ¬a b ¬c = supf0 z
806 ... | tri> ¬a ¬b c = sp1 878 ... | tri> ¬a ¬b c = sp1
807 879
808 sf1=sf0 : {z : Ordinal } → z o≤ px → supf1 z ≡ supf0 z 880 sf1=sf0 : {z : Ordinal } → z o≤ px → supf1 z ≡ supf0 z
809 sf1=sf0 {z} z≤px with trio< z px 881 sf1=sf0 {z} z≤px with trio< z px
810 ... | tri< a ¬b ¬c = refl 882 ... | tri< a ¬b ¬c = refl
817 ... | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (sym b) px<z ) 889 ... | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (sym b) px<z )
818 ... | tri> ¬a ¬b c = refl 890 ... | tri> ¬a ¬b c = refl
819 891
820 asupf1 : {z : Ordinal } → odef A (supf1 z) 892 asupf1 : {z : Ordinal } → odef A (supf1 z)
821 asupf1 {z} with trio< z px 893 asupf1 {z} with trio< z px
822 ... | tri< a ¬b ¬c = ZChain.asupf zc 894 ... | tri< a ¬b ¬c = ZChain.asupf zc
823 ... | tri≈ ¬a b ¬c = ZChain.asupf zc 895 ... | tri≈ ¬a b ¬c = ZChain.asupf zc
824 ... | tri> ¬a ¬b c = MinSUP.asm sup1 896 ... | tri> ¬a ¬b c = MinSUP.asm sup1
825 897
826 supf1-mono : {a b : Ordinal } → a o≤ b → supf1 a o≤ supf1 b 898 supf1-mono : {a b : Ordinal } → a o≤ b → supf1 a o≤ supf1 b
827 supf1-mono {a} {b} a≤b with trio< b px 899 supf1-mono {a} {b} a≤b with trio< b px
828 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (o<→≤ (ordtrans≤-< a≤b a)))) refl ( supf-mono a≤b ) 900 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (o<→≤ (ordtrans≤-< a≤b a)))) refl ( supf-mono a≤b )
829 ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (subst (λ k → a o≤ k) b a≤b))) refl ( supf-mono a≤b ) 901 ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (subst (λ k → a o≤ k) b a≤b))) refl ( supf-mono a≤b )
830 supf1-mono {a} {b} a≤b | tri> ¬a ¬b c with trio< a px 902 supf1-mono {a} {b} a≤b | tri> ¬a ¬b c with trio< a px
831 ... | tri< a<px ¬b ¬c = zc19 where 903 ... | tri< a<px ¬b ¬c = zc19 where
832 zc21 : MinSUP A (UnionCF A f mf ay supf0 a) 904 zc21 : MinSUP A (UnionCF A f mf ay supf0 a)
833 zc21 = ZChain.minsup zc (o<→≤ a<px) 905 zc21 = ZChain.minsup zc (o<→≤ a<px)
834 zc24 : {x₁ : Ordinal} → odef (UnionCF A f mf ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1) 906 zc24 : {x₁ : Ordinal} → odef (UnionCF A f mf ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1)
835 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o<→≤ a<px) ux ) ) 907 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o<→≤ a<px) ux ) )
836 zc19 : supf0 a o≤ sp1 908 zc19 : supf0 a o≤ sp1
837 zc19 = subst (λ k → k o≤ sp1) (sym (ZChain.supf-is-minsup zc (o<→≤ a<px))) ( MinSUP.minsup zc21 (MinSUP.asm sup1) zc24 ) 909 zc19 = subst (λ k → k o≤ sp1) (sym (ZChain.supf-is-minsup zc (o<→≤ a<px))) ( MinSUP.minsup zc21 (MinSUP.asm sup1) zc24 )
838 ... | tri≈ ¬a b ¬c = zc18 where 910 ... | tri≈ ¬a b ¬c = zc18 where
839 zc21 : MinSUP A (UnionCF A f mf ay supf0 a) 911 zc21 : MinSUP A (UnionCF A f mf ay supf0 a)
840 zc21 = ZChain.minsup zc (o≤-refl0 b) 912 zc21 = ZChain.minsup zc (o≤-refl0 b)
841 zc20 : MinSUP.sup zc21 ≡ supf0 a 913 zc20 : MinSUP.sup zc21 ≡ supf0 a
842 zc20 = sym (ZChain.supf-is-minsup zc (o≤-refl0 b)) 914 zc20 = sym (ZChain.supf-is-minsup zc (o≤-refl0 b))
843 zc24 : {x₁ : Ordinal} → odef (UnionCF A f mf ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1) 915 zc24 : {x₁ : Ordinal} → odef (UnionCF A f mf ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1)
844 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o≤-refl0 b) ux ) ) 916 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o≤-refl0 b) ux ) )
845 zc18 : supf0 a o≤ sp1 917 zc18 : supf0 a o≤ sp1
846 zc18 = subst (λ k → k o≤ sp1) zc20( MinSUP.minsup zc21 (MinSUP.asm sup1) zc24 ) 918 zc18 = subst (λ k → k o≤ sp1) zc20( MinSUP.minsup zc21 (MinSUP.asm sup1) zc24 )
847 ... | tri> ¬a ¬b c = o≤-refl 919 ... | tri> ¬a ¬b c = o≤-refl
848 920
849 921
850 fcup : {u z : Ordinal } → FClosure A f (supf1 u) z → u o≤ px → FClosure A f (supf0 u) z 922 fcup : {u z : Ordinal } → FClosure A f (supf1 u) z → u o≤ px → FClosure A f (supf0 u) z
851 fcup {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sf1=sf0 u≤px) fc 923 fcup {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sf1=sf0 u≤px) fc
852 fcpu : {u z : Ordinal } → FClosure A f (supf0 u) z → u o≤ px → FClosure A f (supf1 u) z 924 fcpu : {u z : Ordinal } → FClosure A f (supf0 u) z → u o≤ px → FClosure A f (supf1 u) z
853 fcpu {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sym (sf1=sf0 u≤px)) fc 925 fcpu {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sym (sf1=sf0 u≤px)) fc
854 zc11 : {z : Ordinal} → odef (UnionCF A f mf ay supf1 x) z → odef pchainpx z 926 zc11 : {z : Ordinal} → odef (UnionCF A f mf ay supf1 x) z → odef pchainpx z
927 zc11 {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫
855 zc11 {z} ⟪ az , ch-is-sup u su<sx is-sup fc ⟫ = zc21 fc where 928 zc11 {z} ⟪ az , ch-is-sup u su<sx is-sup fc ⟫ = zc21 fc where
856 u<x : u o< x 929 u<x : u o< x
857 u<x = supf-inject0 supf1-mono su<sx 930 u<x = supf-inject0 supf1-mono su<sx
858 u≤px : u o≤ px 931 u≤px : u o≤ px
859 u≤px = zc-b<x _ u<x 932 u≤px = zc-b<x _ u<x
860 zc21 : {z1 : Ordinal } → FClosure A f (supf1 u) z1 → odef pchainpx z1 933 zc21 : {z1 : Ordinal } → FClosure A f (supf1 u) z1 → odef pchainpx z1
861 zc21 {z1} (fsuc z2 fc ) with zc21 fc 934 zc21 {z1} (fsuc z2 fc ) with zc21 fc
862 ... | case1 ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫ 935 ... | case1 ⟪ ua1 , ch-init fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1) , ch-init (fsuc _ fc₁) ⟫
863 ... | case2 fc = case2 (fsuc _ fc) 936 ... | case1 ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫
937 ... | case2 fc = case2 (fsuc _ fc)
864 zc21 (init asp refl ) with trio< (supf0 u) (supf0 px) | inspect supf1 u 938 zc21 (init asp refl ) with trio< (supf0 u) (supf0 px) | inspect supf1 u
865 ... | tri< a ¬b ¬c | _ = case1 ⟪ asp , ch-is-sup u a record {fcy<sup = zc13 ; order = zc17 939 ... | tri< a ¬b ¬c | _ = case1 ⟪ asp , ch-is-sup u a record {fcy<sup = zc13 ; order = zc17
866 ; supu=u = trans (sym (sf1=sf0 (o<→≤ u<px))) (ChainP.supu=u is-sup) } (init asp0 (sym (sf1=sf0 (o<→≤ u<px))) ) ⟫ where 940 ; supu=u = trans (sym (sf1=sf0 (o<→≤ u<px))) (ChainP.supu=u is-sup) } (init asp0 (sym (sf1=sf0 (o<→≤ u<px))) ) ⟫ where
867 u<px : u o< px 941 u<px : u o< px
868 u<px = ZChain.supf-inject zc a 942 u<px = ZChain.supf-inject zc a
869 asp0 : odef A (supf0 u) 943 asp0 : odef A (supf0 u)
870 asp0 = ZChain.asupf zc 944 asp0 = ZChain.asupf zc
871 zc17 : {s : Ordinal} {z1 : Ordinal} → supf0 s o< supf0 u → 945 zc17 : {s : Ordinal} {z1 : Ordinal} → supf0 s o< supf0 u →
872 FClosure A f (supf0 s) z1 → (z1 ≡ supf0 u) ∨ (z1 << supf0 u) 946 FClosure A f (supf0 s) z1 → (z1 ≡ supf0 u) ∨ (z1 << supf0 u)
873 zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) ((sf1=sf0 u≤px)) ( ChainP.order is-sup 947 zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) ((sf1=sf0 u≤px)) ( ChainP.order is-sup
874 (subst₂ (λ j k → j o< k ) (sym (sf1=sf0 zc18)) (sym (sf1=sf0 u≤px)) ss<spx) (fcpu fc zc18) ) where 948 (subst₂ (λ j k → j o< k ) (sym (sf1=sf0 zc18)) (sym (sf1=sf0 u≤px)) ss<spx) (fcpu fc zc18) ) where
875 zc18 : s o≤ px 949 zc18 : s o≤ px
876 zc18 = ordtrans (ZChain.supf-inject zc ss<spx) u≤px 950 zc18 = ordtrans (ZChain.supf-inject zc ss<spx) u≤px
877 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf0 u) ∨ ( z << supf0 u ) 951 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf0 u) ∨ ( z << supf0 u )
878 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sf1=sf0 (o<→≤ u<px)) ( ChainP.fcy<sup is-sup fc ) 952 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sf1=sf0 (o<→≤ u<px)) ( ChainP.fcy<sup is-sup fc )
879 ... | tri≈ ¬a b ¬c | _ = case2 (init (subst (λ k → odef A k) b (ZChain.asupf zc) ) (sym (trans (sf1=sf0 u≤px) b ))) 953 ... | tri≈ ¬a b ¬c | _ = case2 (init (subst (λ k → odef A k) b (ZChain.asupf zc) ) (sym (trans (sf1=sf0 u≤px) b )))
880 ... | tri> ¬a ¬b c | _ = ⊥-elim ( ¬p<x<op ⟪ ZChain.supf-inject zc c , subst (λ k → u o< k ) (sym (Oprev.oprev=x op)) u<x ⟫ ) 954 ... | tri> ¬a ¬b c | _ = ⊥-elim ( ¬p<x<op ⟪ ZChain.supf-inject zc c , subst (λ k → u o< k ) (sym (Oprev.oprev=x op)) u<x ⟫ )
881 zc12 : {z : Ordinal} → odef pchainpx z → odef (UnionCF A f mf ay supf1 x) z 955 zc12 : {z : Ordinal} → odef pchainpx z → odef (UnionCF A f mf ay supf1 x) z
956 zc12 {z} (case1 ⟪ az , ch-init fc ⟫ ) = ⟪ az , ch-init fc ⟫
882 zc12 {z} (case1 ⟪ az , ch-is-sup u su<sx is-sup fc ⟫ ) = zc21 fc where 957 zc12 {z} (case1 ⟪ az , ch-is-sup u su<sx is-sup fc ⟫ ) = zc21 fc where
883 u<px : u o< px 958 u<px : u o< px
884 u<px = ZChain.supf-inject zc su<sx 959 u<px = ZChain.supf-inject zc su<sx
885 u<x : u o< x 960 u<x : u o< x
886 u<x = ordtrans u<px px<x 961 u<x = ordtrans u<px px<x
887 u≤px : u o≤ px 962 u≤px : u o≤ px
888 u≤px = o<→≤ u<px 963 u≤px = o<→≤ u<px
889 s1u<s1x : supf1 u o< supf1 x 964 s1u<s1x : supf1 u o< supf1 x
890 s1u<s1x = ordtrans<-≤ (subst₂ (λ j k → j o< k ) (sym (sf1=sf0 u≤px )) (sym (sf1=sf0 o≤-refl)) su<sx) (supf1-mono (o<→≤ px<x) ) 965 s1u<s1x = ordtrans<-≤ (subst₂ (λ j k → j o< k ) (sym (sf1=sf0 u≤px )) (sym (sf1=sf0 o≤-refl)) su<sx) (supf1-mono (o<→≤ px<x) )
891 zc21 : {z1 : Ordinal } → FClosure A f (supf0 u) z1 → odef (UnionCF A f mf ay supf1 x) z1 966 zc21 : {z1 : Ordinal } → FClosure A f (supf0 u) z1 → odef (UnionCF A f mf ay supf1 x) z1
892 zc21 {z1} (fsuc z2 fc ) with zc21 fc 967 zc21 {z1} (fsuc z2 fc ) with zc21 fc
893 ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫ 968 ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-init (fsuc _ fc₁) ⟫
969 ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫
894 zc21 (init asp refl ) with trio< u px | inspect supf1 u 970 zc21 (init asp refl ) with trio< u px | inspect supf1 u
895 ... | tri< a ¬b ¬c | _ = ⟪ asp , ch-is-sup u 971 ... | tri< a ¬b ¬c | _ = ⟪ asp , ch-is-sup u
896 s1u<s1x 972 s1u<s1x
897 record {fcy<sup = zc13 ; order = zc17 ; supu=u = trans (sf1=sf0 u≤px ) (ChainP.supu=u is-sup) } zc14 ⟫ where 973 record {fcy<sup = zc13 ; order = zc17 ; supu=u = trans (sf1=sf0 u≤px ) (ChainP.supu=u is-sup) } zc14 ⟫ where
898 zc17 : {s : Ordinal} {z1 : Ordinal} → supf1 s o< supf1 u → 974 zc17 : {s : Ordinal} {z1 : Ordinal} → supf1 s o< supf1 u →
899 FClosure A f (supf1 s) z1 → (z1 ≡ supf1 u) ∨ (z1 << supf1 u) 975 FClosure A f (supf1 s) z1 → (z1 ≡ supf1 u) ∨ (z1 << supf1 u)
900 zc17 {s} {z1} ss<su fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) (sym (sf1=sf0 (o<→≤ u<px))) ( ChainP.order is-sup 976 zc17 {s} {z1} ss<su fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) (sym (sf1=sf0 (o<→≤ u<px))) ( ChainP.order is-sup
901 (subst₂ (λ j k → j o< k ) (sf1=sf0 s≤px) (sf1=sf0 (o<→≤ u<px)) ss<su) (fcup fc s≤px) ) where 977 (subst₂ (λ j k → j o< k ) (sf1=sf0 s≤px) (sf1=sf0 (o<→≤ u<px)) ss<su) (fcup fc s≤px) ) where
902 s≤px : s o≤ px -- ss<su : supf1 s o< supf1 u 978 s≤px : s o≤ px -- ss<su : supf1 s o< supf1 u
903 s≤px = ordtrans ( supf-inject0 supf1-mono ss<su ) (o<→≤ u<px) 979 s≤px = ordtrans ( supf-inject0 supf1-mono ss<su ) (o<→≤ u<px)
904 zc14 : FClosure A f (supf1 u) (supf0 u) 980 zc14 : FClosure A f (supf1 u) (supf0 u)
905 zc14 = init (subst (λ k → odef A k ) (sym (sf1=sf0 u≤px)) asp) (sf1=sf0 u≤px) 981 zc14 = init (subst (λ k → odef A k ) (sym (sf1=sf0 u≤px)) asp) (sf1=sf0 u≤px)
906 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 u) ∨ ( z << supf1 u ) 982 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 u) ∨ ( z << supf1 u )
907 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 u≤px )) ( ChainP.fcy<sup is-sup fc ) 983 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 u≤px )) ( ChainP.fcy<sup is-sup fc )
908 ... | tri≈ ¬a b ¬c | _ = ⟪ asp , ch-is-sup px (subst (λ k → supf1 k o< supf1 x) b s1u<s1x) record { fcy<sup = zc13 984 ... | tri≈ ¬a b ¬c | _ = ⟪ asp , ch-is-sup px (subst (λ k → supf1 k o< supf1 x) b s1u<s1x) record { fcy<sup = zc13
909 ; order = zc17 ; supu=u = zc18 } (init asupf1 (trans (sf1=sf0 o≤-refl ) (cong supf0 (sym b))) ) ⟫ where 985 ; order = zc17 ; supu=u = zc18 } (init asupf1 (trans (sf1=sf0 o≤-refl ) (cong supf0 (sym b))) ) ⟫ where
910 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 px) ∨ ( z << supf1 px ) 986 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 px) ∨ ( z << supf1 px )
911 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ (z << k)) (trans (cong supf0 b) (sym (sf1=sf0 o≤-refl))) (ChainP.fcy<sup is-sup fc ) 987 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ (z << k)) (trans (cong supf0 b) (sym (sf1=sf0 o≤-refl))) (ChainP.fcy<sup is-sup fc )
912 zc18 : supf1 px ≡ px 988 zc18 : supf1 px ≡ px
913 zc18 = begin 989 zc18 = begin
929 --- supf0 px o< sp1 1005 --- supf0 px o< sp1
930 zc20 : (supf0 px ≡ px ) ∨ ( supf0 px o< px ) 1006 zc20 : (supf0 px ≡ px ) ∨ ( supf0 px o< px )
931 zc20 = ? 1007 zc20 = ?
932 zc21 : {z1 : Ordinal } → FClosure A f (supf0 px) z1 → odef (UnionCF A f mf ay supf1 x) z1 1008 zc21 : {z1 : Ordinal } → FClosure A f (supf0 px) z1 → odef (UnionCF A f mf ay supf1 x) z1
933 zc21 {z1} (fsuc z2 fc ) with zc21 fc 1009 zc21 {z1} (fsuc z2 fc ) with zc21 fc
934 ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫ 1010 ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-init (fsuc _ fc₁) ⟫
1011 ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫
935 zc21 (init asp refl ) with zc20 1012 zc21 (init asp refl ) with zc20
936 ... | case1 sfpx=px = ⟪ asp , ch-is-sup px zc18 1013 ... | case1 sfpx=px = ⟪ asp , ch-is-sup px zc18
937 record {fcy<sup = zc13 ; order = zc17 ; supu=u = zc15 } zc14 ⟫ where 1014 record {fcy<sup = zc13 ; order = zc17 ; supu=u = zc15 } zc14 ⟫ where
938 zc15 : supf1 px ≡ px 1015 zc15 : supf1 px ≡ px
939 zc15 = trans (sf1=sf0 o≤-refl ) (sfpx=px) 1016 zc15 = trans (sf1=sf0 o≤-refl ) (sfpx=px)
940 zc18 : supf1 px o< supf1 x 1017 zc18 : supf1 px o< supf1 x
941 zc18 = ? 1018 zc18 = ?
943 zc14 = init (subst (λ k → odef A k) (sym (sf1=sf0 o≤-refl)) asp) (sf1=sf0 o≤-refl) 1020 zc14 = init (subst (λ k → odef A k) (sym (sf1=sf0 o≤-refl)) asp) (sf1=sf0 o≤-refl)
944 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 px ) ∨ ( z << supf1 px ) 1021 zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 px ) ∨ ( z << supf1 px )
945 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 o≤-refl)) ( ZChain.fcy<sup zc o≤-refl fc ) 1022 zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 o≤-refl)) ( ZChain.fcy<sup zc o≤-refl fc )
946 zc17 : {s : Ordinal} {z1 : Ordinal} → supf1 s o< supf1 px → 1023 zc17 : {s : Ordinal} {z1 : Ordinal} → supf1 s o< supf1 px →
947 FClosure A f (supf1 s) z1 → (z1 ≡ supf1 px) ∨ (z1 << supf1 px) 1024 FClosure A f (supf1 s) z1 → (z1 ≡ supf1 px) ∨ (z1 << supf1 px)
948 zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k )) mins-is-spx 1025 zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k )) mins-is-spx
949 (MinSUP.x≤sup mins (csupf17 (fcup fc (o<→≤ s<px) )) ) where 1026 (MinSUP.x≤sup mins (csupf17 (fcup fc (o<→≤ s<px) )) ) where
950 mins : MinSUP A (UnionCF A f mf ay supf0 px) 1027 mins : MinSUP A (UnionCF A f mf ay supf0 px)
951 mins = ZChain.minsup zc o≤-refl 1028 mins = ZChain.minsup zc o≤-refl
952 mins-is-spx : MinSUP.sup mins ≡ supf1 px 1029 mins-is-spx : MinSUP.sup mins ≡ supf1 px
953 mins-is-spx = trans (sym ( ZChain.supf-is-minsup zc o≤-refl ) ) (sym (sf1=sf0 o≤-refl )) 1030 mins-is-spx = trans (sym ( ZChain.supf-is-minsup zc o≤-refl ) ) (sym (sf1=sf0 o≤-refl ))
954 s<px : s o< px 1031 s<px : s o< px
955 s<px = supf-inject0 supf1-mono ss<spx 1032 s<px = supf-inject0 supf1-mono ss<spx
956 sf<px : supf0 s o< px 1033 sf<px : supf0 s o< px
957 sf<px = subst₂ (λ j k → j o< k ) (sf1=sf0 (o<→≤ s<px)) (trans (sf1=sf0 o≤-refl) (sfpx=px)) ss<spx 1034 sf<px = subst₂ (λ j k → j o< k ) (sf1=sf0 (o<→≤ s<px)) (trans (sf1=sf0 o≤-refl) (sfpx=px)) ss<spx
958 csupf17 : {z1 : Ordinal } → FClosure A f (supf0 s) z1 → odef (UnionCF A f mf ay supf0 px) z1 1035 csupf17 : {z1 : Ordinal } → FClosure A f (supf0 s) z1 → odef (UnionCF A f mf ay supf0 px) z1
959 csupf17 (init as refl ) = ZChain.csupf zc sf<px 1036 csupf17 (init as refl ) = ZChain.csupf zc sf<px
960 csupf17 (fsuc x fc) = ZChain.f-next zc (csupf17 fc) 1037 csupf17 (fsuc x fc) = ZChain.f-next zc (csupf17 fc)
961 1038
962 ... | case2 sfp<px with ZChain.csupf zc sfp<px -- odef (UnionCF A f mf ay supf0 px) (supf0 px) 1039 ... | case2 sfp<px with ZChain.csupf zc sfp<px -- odef (UnionCF A f mf ay supf0 px) (supf0 px)
963 ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ ua1 , ch-is-sup u zc18 1040 ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ ua1 , ch-init fc₁ ⟫
1041 ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ ua1 , ch-is-sup u zc18
964 record { fcy<sup = z10 ; order = z11 ; supu=u = z12 } (fcpu fc₁ ? ) ⟫ where 1042 record { fcy<sup = z10 ; order = z11 ; supu=u = z12 } (fcpu fc₁ ? ) ⟫ where
965 z10 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 u) ∨ ( z << supf1 u ) 1043 z10 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 u) ∨ ( z << supf1 u )
966 z10 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? )) (ChainP.fcy<sup is-sup fc) 1044 z10 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? )) (ChainP.fcy<sup is-sup fc)
967 z11 : {s z1 : Ordinal} → (lt : supf1 s o< supf1 u ) → FClosure A f (supf1 s ) z1 1045 z11 : {s z1 : Ordinal} → (lt : supf1 s o< supf1 u ) → FClosure A f (supf1 s ) z1
968 → (z1 ≡ supf1 u ) ∨ ( z1 << supf1 u ) 1046 → (z1 ≡ supf1 u ) ∨ ( z1 << supf1 u )
969 z11 {s} {z} lt fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? )) 1047 z11 {s} {z} lt fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? ))
970 (ChainP.order is-sup lt0 (fcup fc s≤px )) where 1048 (ChainP.order is-sup lt0 (fcup fc s≤px )) where
971 s<u : s o< u 1049 s<u : s o< u
972 s<u = supf-inject0 supf1-mono lt 1050 s<u = supf-inject0 supf1-mono lt
973 s≤px : s o≤ px 1051 s≤px : s o≤ px
974 s≤px = ordtrans s<u ? -- (o<→≤ u<x) 1052 s≤px = ordtrans s<u ? -- (o<→≤ u<x)
982 1060
983 1061
984 record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where 1062 record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where
985 field 1063 field
986 tsup : MinSUP A (UnionCF A f mf ay supf1 z) 1064 tsup : MinSUP A (UnionCF A f mf ay supf1 z)
987 tsup=sup : supf1 z ≡ MinSUP.sup tsup 1065 tsup=sup : supf1 z ≡ MinSUP.sup tsup
988 1066
989 sup : {z : Ordinal} → (z≤x : z o≤ x ) → STMP z≤x 1067 sup : {z : Ordinal} → (z≤x : z o≤ x ) → STMP z≤x
990 sup {z} z≤x with trio< z px 1068 sup {z} z≤x with trio< z px
991 ... | tri< a ¬b ¬c = record { tsup = record { sup = MinSUP.sup m ; asm = MinSUP.asm m 1069 ... | tri< a ¬b ¬c = record { tsup = record { sup = MinSUP.sup m ; asm = MinSUP.asm m
992 ; x≤sup = ms00 ; minsup = ms01 } ; tsup=sup = trans (sf1=sf0 (o<→≤ a) ) (ZChain.supf-is-minsup zc (o<→≤ a)) } where 1070 ; x≤sup = ms00 ; minsup = ms01 } ; tsup=sup = trans (sf1=sf0 (o<→≤ a) ) (ZChain.supf-is-minsup zc (o<→≤ a)) } where
993 m = ZChain.minsup zc (o<→≤ a) 1071 m = ZChain.minsup zc (o<→≤ a)
994 ms00 : {x : Ordinal} → odef (UnionCF A f mf ay supf1 z) x → (x ≡ MinSUP.sup m) ∨ (x << MinSUP.sup m) 1072 ms00 : {x : Ordinal} → odef (UnionCF A f mf ay supf1 z) x → (x ≡ MinSUP.sup m) ∨ (x << MinSUP.sup m)
995 ms00 {x} ux = MinSUP.x≤sup m ? 1073 ms00 {x} ux = MinSUP.x≤sup m ?
996 ms01 : {sup2 : Ordinal} → odef A sup2 → ({x : Ordinal} → 1074 ms01 : {sup2 : Ordinal} → odef A sup2 → ({x : Ordinal} →
997 odef (UnionCF A f mf ay supf1 z) x → (x ≡ sup2) ∨ (x << sup2)) → MinSUP.sup m o≤ sup2 1075 odef (UnionCF A f mf ay supf1 z) x → (x ≡ sup2) ∨ (x << sup2)) → MinSUP.sup m o≤ sup2
998 ms01 {sup2} us P = MinSUP.minsup m ? ? 1076 ms01 {sup2} us P = MinSUP.minsup m ? ?
999 ... | tri≈ ¬a b ¬c = record { tsup = record { sup = MinSUP.sup m ; asm = MinSUP.asm m 1077 ... | tri≈ ¬a b ¬c = record { tsup = record { sup = MinSUP.sup m ; asm = MinSUP.asm m
1000 ; x≤sup = ms00 ; minsup = ms01 } ; tsup=sup = trans (sf1=sf0 (o≤-refl0 b) ) (ZChain.supf-is-minsup zc (o≤-refl0 b)) } where 1078 ; x≤sup = ms00 ; minsup = ms01 } ; tsup=sup = trans (sf1=sf0 (o≤-refl0 b) ) (ZChain.supf-is-minsup zc (o≤-refl0 b)) } where
1001 m = ZChain.minsup zc (o≤-refl0 b) 1079 m = ZChain.minsup zc (o≤-refl0 b)
1002 ms00 : {x : Ordinal} → odef (UnionCF A f mf ay supf1 z) x → (x ≡ MinSUP.sup m) ∨ (x << MinSUP.sup m) 1080 ms00 : {x : Ordinal} → odef (UnionCF A f mf ay supf1 z) x → (x ≡ MinSUP.sup m) ∨ (x << MinSUP.sup m)
1003 ms00 {x} ux = MinSUP.x≤sup m ? 1081 ms00 {x} ux = MinSUP.x≤sup m ?
1004 ms01 : {sup2 : Ordinal} → odef A sup2 → ({x : Ordinal} → 1082 ms01 : {sup2 : Ordinal} → odef A sup2 → ({x : Ordinal} →
1011 ms00 {x} ux = MinSUP.x≤sup m ? 1089 ms00 {x} ux = MinSUP.x≤sup m ?
1012 ms01 : {sup2 : Ordinal} → odef A sup2 → ({x : Ordinal} → 1090 ms01 : {sup2 : Ordinal} → odef A sup2 → ({x : Ordinal} →
1013 odef (UnionCF A f mf ay supf1 z) x → (x ≡ sup2) ∨ (x << sup2)) → MinSUP.sup m o≤ sup2 1091 odef (UnionCF A f mf ay supf1 z) x → (x ≡ sup2) ∨ (x << sup2)) → MinSUP.sup m o≤ sup2
1014 ms01 {sup2} us P = MinSUP.minsup m ? ? 1092 ms01 {sup2} us P = MinSUP.minsup m ? ?
1015 1093
1016 csupf1 : {z1 : Ordinal } → supf1 z1 o< x → odef (UnionCF A f mf ay supf1 x) (supf1 z1) 1094 csupf1 : {z1 : Ordinal } → supf1 z1 o< x → odef (UnionCF A f mf ay supf1 x) (supf1 z1)
1017 csupf1 {z1} sz1<x = csupf2 where 1095 csupf1 {z1} sz1<x = csupf2 where
1018 -- z1 o< px → supf1 z1 ≡ supf0 z1 1096 -- z1 o< px → supf1 z1 ≡ supf0 z1
1019 -- z1 ≡ px , supf0 px o< px .. px o< z1, x o≤ z1 ... supf1 z1 ≡ sp1 1097 -- z1 ≡ px , supf0 px o< px .. px o< z1, x o≤ z1 ... supf1 z1 ≡ sp1
1020 -- z1 ≡ px , supf0 px ≡ px 1098 -- z1 ≡ px , supf0 px ≡ px
1021 psz1≤px : supf1 z1 o≤ px 1099 psz1≤px : supf1 z1 o≤ px
1022 psz1≤px = subst (λ k → supf1 z1 o< k ) (sym opx=x) sz1<x 1100 psz1≤px = subst (λ k → supf1 z1 o< k ) (sym opx=x) sz1<x
1023 csupf2 : odef (UnionCF A f mf ay supf1 x) (supf1 z1) 1101 csupf2 : odef (UnionCF A f mf ay supf1 x) (supf1 z1)
1024 csupf2 with trio< z1 px | inspect supf1 z1 1102 csupf2 with trio< z1 px | inspect supf1 z1
1025 csupf2 | tri< a ¬b ¬c | record { eq = eq1 } with osuc-≡< psz1≤px 1103 csupf2 | tri< a ¬b ¬c | record { eq = eq1 } with osuc-≡< psz1≤px
1026 ... | case2 lt = zc12 (case1 cs03) where 1104 ... | case2 lt = zc12 (case1 cs03) where
1027 cs03 : odef (UnionCF A f mf ay supf0 px) (supf0 z1) 1105 cs03 : odef (UnionCF A f mf ay supf0 px) (supf0 z1)
1028 cs03 = ZChain.csupf zc (subst (λ k → k o< px) (sf1=sf0 (o<→≤ a)) lt ) 1106 cs03 = ZChain.csupf zc (subst (λ k → k o< px) (sf1=sf0 (o<→≤ a)) lt )
1029 ... | case1 sfz=px with osuc-≡< ( supf1-mono (o<→≤ a) ) 1107 ... | case1 sfz=px with osuc-≡< ( supf1-mono (o<→≤ a) )
1030 ... | case1 sfz=sfpx = zc12 (case2 (init (ZChain.asupf zc) cs04 )) where 1108 ... | case1 sfz=sfpx = zc12 (case2 (init (ZChain.asupf zc) cs04 )) where
1031 supu=u : supf1 (supf1 z1) ≡ supf1 z1 1109 supu=u : supf1 (supf1 z1) ≡ supf1 z1
1032 supu=u = trans (cong supf1 sfz=px) (sym sfz=sfpx) 1110 supu=u = trans (cong supf1 sfz=px) (sym sfz=sfpx)
1033 cs04 : supf0 px ≡ supf0 z1 1111 cs04 : supf0 px ≡ supf0 z1
1034 cs04 = begin 1112 cs04 = begin
1035 supf0 px ≡⟨ sym (sf1=sf0 o≤-refl) ⟩ 1113 supf0 px ≡⟨ sym (sf1=sf0 o≤-refl) ⟩
1036 supf1 px ≡⟨ sym sfz=sfpx ⟩ 1114 supf1 px ≡⟨ sym sfz=sfpx ⟩
1037 supf1 z1 ≡⟨ sf1=sf0 (o<→≤ a) ⟩ 1115 supf1 z1 ≡⟨ sf1=sf0 (o<→≤ a) ⟩
1038 supf0 z1 ∎ where open ≡-Reasoning 1116 supf0 z1 ∎ where open ≡-Reasoning
1041 cs05 : px o< supf0 px 1119 cs05 : px o< supf0 px
1042 cs05 = subst₂ ( λ j k → j o< k ) sfz=px (sf1=sf0 o≤-refl ) sfz<sfpx 1120 cs05 = subst₂ ( λ j k → j o< k ) sfz=px (sf1=sf0 o≤-refl ) sfz<sfpx
1043 cs06 : supf0 px o< osuc px 1121 cs06 : supf0 px o< osuc px
1044 cs06 = subst (λ k → supf0 px o< k ) (sym opx=x) ? 1122 cs06 = subst (λ k → supf0 px o< k ) (sym opx=x) ?
1045 csupf2 | tri≈ ¬a b ¬c | record { eq = eq1 } = zc12 (case2 (init (ZChain.asupf zc) (cong supf0 (sym b)))) 1123 csupf2 | tri≈ ¬a b ¬c | record { eq = eq1 } = zc12 (case2 (init (ZChain.asupf zc) (cong supf0 (sym b))))
1046 csupf2 | tri> ¬a ¬b px<z1 | record { eq = eq1 } = ? 1124 csupf2 | tri> ¬a ¬b px<z1 | record { eq = eq1 } = ?
1047 -- ⊥-elim ( ¬p<x<op ⟪ px<z1 , subst (λ k → z1 o< k) (sym opx=x) z1<x ⟫ ) 1125 -- ⊥-elim ( ¬p<x<op ⟪ px<z1 , subst (λ k → z1 o< k) (sym opx=x) z1<x ⟫ )
1048 1126
1049 1127
1050 zc41 | (case1 sfp=sp1 ) = record { supf = supf0 ; sup=u = ? ; asupf = ? ; supf-mono = ? ; supf-< = ? 1128 zc41 | (case1 sfp=sp1 ) = record { supf = supf0 ; sup=u = ? ; asupf = ? ; supf-mono = ? ; supf-< = ?
1051 ; supfmax = ? ; minsup = ? ; supf-is-minsup = ? ; csupf = ? } where 1129 ; supfmax = ? ; minsup = ? ; supf-is-minsup = ? ; csupf = ? } where
1052 1130
1053 -- supf0 px not is included by the chain 1131 -- supf0 px not is included by the chain
1054 -- supf1 x ≡ supf0 px because of supfmax 1132 -- supf1 x ≡ supf0 px because of supfmax
1055 1133
1056 supf1 : Ordinal → Ordinal 1134 supf1 : Ordinal → Ordinal
1057 supf1 z with trio< z px 1135 supf1 z with trio< z px
1058 ... | tri< a ¬b ¬c = supf0 z 1136 ... | tri< a ¬b ¬c = supf0 z
1059 ... | tri≈ ¬a b ¬c = supf0 px 1137 ... | tri≈ ¬a b ¬c = supf0 px
1060 ... | tri> ¬a ¬b c = supf0 px 1138 ... | tri> ¬a ¬b c = supf0 px
1061 1139
1062 sf1=sf0 : {z : Ordinal } → z o< px → supf1 z ≡ supf0 z 1140 sf1=sf0 : {z : Ordinal } → z o< px → supf1 z ≡ supf0 z
1067 1145
1068 zc17 : {z : Ordinal } → supf0 z o≤ supf0 px 1146 zc17 : {z : Ordinal } → supf0 z o≤ supf0 px
1069 zc17 = ? -- px o< z, px o< supf0 px 1147 zc17 = ? -- px o< z, px o< supf0 px
1070 1148
1071 supf-mono1 : {z w : Ordinal } → z o≤ w → supf1 z o≤ supf1 w 1149 supf-mono1 : {z w : Ordinal } → z o≤ w → supf1 z o≤ supf1 w
1072 supf-mono1 {z} {w} z≤w with trio< w px 1150 supf-mono1 {z} {w} z≤w with trio< w px
1073 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (ordtrans≤-< z≤w a))) refl ( supf-mono z≤w ) 1151 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (ordtrans≤-< z≤w a))) refl ( supf-mono z≤w )
1074 ... | tri≈ ¬a refl ¬c with trio< z px 1152 ... | tri≈ ¬a refl ¬c with trio< z px
1075 ... | tri< a ¬b ¬c = zc17 1153 ... | tri< a ¬b ¬c = zc17
1076 ... | tri≈ ¬a refl ¬c = o≤-refl 1154 ... | tri≈ ¬a refl ¬c = o≤-refl
1077 ... | tri> ¬a ¬b c = o≤-refl 1155 ... | tri> ¬a ¬b c = o≤-refl
1082 1160
1083 pchain1 : HOD 1161 pchain1 : HOD
1084 pchain1 = UnionCF A f mf ay supf1 x 1162 pchain1 = UnionCF A f mf ay supf1 x
1085 1163
1086 zc10 : {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchain1) z 1164 zc10 : {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchain1) z
1165 zc10 {z} ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1087 zc10 {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ⟪ az , ch-is-sup u1 ? ? ? ⟫ 1166 zc10 {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ⟪ az , ch-is-sup u1 ? ? ? ⟫
1088 1167
1089 zc111 : {z : Ordinal} → z o< px → OD.def (od pchain1) z → OD.def (od pchain) z 1168 zc111 : {z : Ordinal} → z o< px → OD.def (od pchain1) z → OD.def (od pchain) z
1169 zc111 {z} z<px ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1090 zc111 {z} z<px ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ⟪ az , ch-is-sup u1 ? ? ? ⟫ 1170 zc111 {z} z<px ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ⟪ az , ch-is-sup u1 ? ? ? ⟫
1091 1171
1092 zc11 : (¬ xSUP (UnionCF A f mf ay supf0 px) f x ) ∨ (HasPrev A pchain f x ) 1172 zc11 : (¬ xSUP (UnionCF A f mf ay supf0 px) f x ) ∨ (HasPrev A pchain f x )
1093 → {z : Ordinal} → OD.def (od pchain1) z → OD.def (od pchain) z ∨ ( (supf0 px ≡ px) ∧ FClosure A f px z ) 1173 → {z : Ordinal} → OD.def (od pchain1) z → OD.def (od pchain) z ∨ ( (supf0 px ≡ px) ∧ FClosure A f px z )
1094 zc11 P {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ with trio< u1 px 1174 zc11 P {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫
1095 ... | tri< u1<px ¬b ¬c = case1 ⟪ az , ch-is-sup u1 ? ? fc ⟫ 1175 zc11 P {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ with trio< u1 px
1176 ... | tri< u1<px ¬b ¬c = case1 ⟪ az , ch-is-sup u1 ? ? fc ⟫
1096 ... | tri≈ ¬a eq ¬c = case2 ⟪ subst (λ k → supf0 k ≡ k) eq s1u=u , subst (λ k → FClosure A f k z) zc12 ? ⟫ where 1177 ... | tri≈ ¬a eq ¬c = case2 ⟪ subst (λ k → supf0 k ≡ k) eq s1u=u , subst (λ k → FClosure A f k z) zc12 ? ⟫ where
1097 s1u=u : supf0 u1 ≡ u1 1178 s1u=u : supf0 u1 ≡ u1
1098 s1u=u = ? -- ChainP.supu=u u1-is-sup 1179 s1u=u = ? -- ChainP.supu=u u1-is-sup
1099 zc12 : supf0 u1 ≡ px 1180 zc12 : supf0 u1 ≡ px
1100 zc12 = trans s1u=u eq 1181 zc12 = trans s1u=u eq
1101 zc11 (case1 ¬sp=x) {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ | tri> ¬a ¬b px<u = ⊥-elim (¬sp=x zcsup) where 1182 zc11 (case1 ¬sp=x) {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ | tri> ¬a ¬b px<u = ⊥-elim (¬sp=x zcsup) where
1102 eq : u1 ≡ x 1183 eq : u1 ≡ x
1103 eq with trio< u1 x 1184 eq with trio< u1 x
1104 ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<u , subst (λ k → u1 o< k ) (sym (Oprev.oprev=x op )) a ⟫ ) 1185 ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<u , subst (λ k → u1 o< k ) (sym (Oprev.oprev=x op )) a ⟫ )
1105 ... | tri≈ ¬a b ¬c = b 1186 ... | tri≈ ¬a b ¬c = b
1106 ... | tri> ¬a ¬b c = ⊥-elim ( o<> u1<x ? ) 1187 ... | tri> ¬a ¬b c = ⊥-elim ( o<> u1<x ? )
1107 s1u=x : supf0 u1 ≡ x 1188 s1u=x : supf0 u1 ≡ x
1108 s1u=x = trans ? eq 1189 s1u=x = trans ? eq
1109 zc13 : osuc px o< osuc u1 1190 zc13 : osuc px o< osuc u1
1110 zc13 = o≤-refl0 ( trans (Oprev.oprev=x op) (sym eq ) ) 1191 zc13 = o≤-refl0 ( trans (Oprev.oprev=x op) (sym eq ) )
1111 x≤sup : {w : Ordinal} → odef (UnionCF A f mf ay supf0 px) w → (w ≡ x) ∨ (w << x) 1192 x≤sup : {w : Ordinal} → odef (UnionCF A f mf ay supf0 px) w → (w ≡ x) ∨ (w << x)
1193 x≤sup {w} ⟪ az , ch-init {w} fc ⟫ = subst (λ k → (w ≡ k) ∨ (w << k)) s1u=x ?
1112 x≤sup {w} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ with osuc-≡< ( supf-mono (ordtrans (o<→≤ u<x) ? )) 1194 x≤sup {w} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ with osuc-≡< ( supf-mono (ordtrans (o<→≤ u<x) ? ))
1113 ... | case1 eq1 = ⊥-elim ( o<¬≡ zc14 ? ) where 1195 ... | case1 eq1 = ⊥-elim ( o<¬≡ zc14 ? ) where
1114 zc14 : u ≡ osuc px 1196 zc14 : u ≡ osuc px
1115 zc14 = begin 1197 zc14 = begin
1116 u ≡⟨ sym ( ChainP.supu=u is-sup) ⟩ 1198 u ≡⟨ sym ( ChainP.supu=u is-sup) ⟩
1117 supf0 u ≡⟨ ? ⟩ 1199 supf0 u ≡⟨ ? ⟩
1118 supf0 u1 ≡⟨ s1u=x ⟩ 1200 supf0 u1 ≡⟨ s1u=x ⟩
1119 x ≡⟨ sym (Oprev.oprev=x op) ⟩ 1201 x ≡⟨ sym (Oprev.oprev=x op) ⟩
1120 osuc px ∎ where open ≡-Reasoning 1202 osuc px ∎ where open ≡-Reasoning
1121 ... | case2 lt = subst (λ k → (w ≡ k) ∨ (w << k)) s1u=x ? 1203 ... | case2 lt = subst (λ k → (w ≡ k) ∨ (w << k)) s1u=x ?
1122 zc12 : supf0 x ≡ u1 1204 zc12 : supf0 x ≡ u1
1123 zc12 = subst (λ k → supf0 k ≡ u1) eq ? 1205 zc12 = subst (λ k → supf0 k ≡ u1) eq ?
1124 zcsup : xSUP (UnionCF A f mf ay supf0 px) f x 1206 zcsup : xSUP (UnionCF A f mf ay supf0 px) f x
1125 zcsup = record { ax = subst (λ k → odef A k) (trans zc12 eq) (ZChain.asupf zc) 1207 zcsup = record { ax = subst (λ k → odef A k) (trans zc12 eq) (ZChain.asupf zc)
1126 ; is-sup = record { x≤sup = x≤sup ; minsup = ? } } 1208 ; is-sup = record { x≤sup = x≤sup ; minsup = ? } }
1127 zc11 (case2 hp) {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ | tri> ¬a ¬b px<u = case1 ? where 1209 zc11 (case2 hp) {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ | tri> ¬a ¬b px<u = case1 ? where
1128 eq : u1 ≡ x 1210 eq : u1 ≡ x
1129 eq with trio< u1 x 1211 eq with trio< u1 x
1130 ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<u , subst (λ k → u1 o< k ) (sym (Oprev.oprev=x op )) a ⟫ ) 1212 ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<u , subst (λ k → u1 o< k ) (sym (Oprev.oprev=x op )) a ⟫ )
1131 ... | tri≈ ¬a b ¬c = b 1213 ... | tri≈ ¬a b ¬c = b
1132 ... | tri> ¬a ¬b c = ⊥-elim ( o<> u1<x ? ) 1214 ... | tri> ¬a ¬b c = ⊥-elim ( o<> u1<x ? )
1133 zc20 : {z : Ordinal} → FClosure A f (supf0 u1) z → OD.def (od pchain) z 1215 zc20 : {z : Ordinal} → FClosure A f (supf0 u1) z → OD.def (od pchain) z
1143 zc20 {.(f w)} (fsuc w fc) = ZChain.f-next zc (zc20 fc) 1225 zc20 {.(f w)} (fsuc w fc) = ZChain.f-next zc (zc20 fc)
1144 1226
1145 record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where 1227 record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where
1146 field 1228 field
1147 tsup : MinSUP A (UnionCF A f mf ay supf1 z) 1229 tsup : MinSUP A (UnionCF A f mf ay supf1 z)
1148 tsup=sup : supf1 z ≡ MinSUP.sup tsup 1230 tsup=sup : supf1 z ≡ MinSUP.sup tsup
1149 1231
1150 sup : {z : Ordinal} → (z≤x : z o≤ x ) → STMP z≤x 1232 sup : {z : Ordinal} → (z≤x : z o≤ x ) → STMP z≤x
1151 sup {z} z≤x with trio< z px 1233 sup {z} z≤x with trio< z px
1152 ... | tri< a ¬b ¬c = ? -- jrecord { tsup = ZChain.minsup zc (o<→≤ a) ; tsup=sup = ZChain.supf-is-minsup zc (o<→≤ a) } 1234 ... | tri< a ¬b ¬c = ? -- jrecord { tsup = ZChain.minsup zc (o<→≤ a) ; tsup=sup = ZChain.supf-is-minsup zc (o<→≤ a) }
1153 ... | tri≈ ¬a b ¬c = ? -- record { tsup = ZChain.minsup zc (o≤-refl0 b) ; tsup=sup = ZChain.supf-is-minsup zc (o≤-refl0 b) } 1235 ... | tri≈ ¬a b ¬c = ? -- record { tsup = ZChain.minsup zc (o≤-refl0 b) ; tsup=sup = ZChain.supf-is-minsup zc (o≤-refl0 b) }
1154 ... | tri> ¬a ¬b px<z = zc35 where 1236 ... | tri> ¬a ¬b px<z = zc35 where
1155 zc30 : z ≡ x 1237 zc30 : z ≡ x
1156 zc30 with osuc-≡< z≤x 1238 zc30 with osuc-≡< z≤x
1157 ... | case1 eq = eq 1239 ... | case1 eq = eq
1158 ... | case2 z<x = ⊥-elim (¬p<x<op ⟪ px<z , subst (λ k → z o< k ) (sym (Oprev.oprev=x op)) z<x ⟫ ) 1240 ... | case2 z<x = ⊥-elim (¬p<x<op ⟪ px<z , subst (λ k → z o< k ) (sym (Oprev.oprev=x op)) z<x ⟫ )
1159 zc32 = ZChain.sup zc o≤-refl 1241 zc32 = ZChain.sup zc o≤-refl
1160 zc34 : ¬ (supf0 px ≡ px) → {w : HOD} → UnionCF A f mf ay supf0 z ∋ w → (w ≡ SUP.sup zc32) ∨ (w < SUP.sup zc32) 1242 zc34 : ¬ (supf0 px ≡ px) → {w : HOD} → UnionCF A f mf ay supf0 z ∋ w → (w ≡ SUP.sup zc32) ∨ (w < SUP.sup zc32)
1161 zc34 ne {w} lt with zc11 ? ⟪ proj1 lt , ? ⟫ 1243 zc34 ne {w} lt with zc11 ? ⟪ proj1 lt , ? ⟫
1162 ... | case1 lt = SUP.x≤sup zc32 lt 1244 ... | case1 lt = SUP.x≤sup zc32 lt
1163 ... | case2 ⟪ spx=px , fc ⟫ = ⊥-elim ( ne spx=px ) 1245 ... | case2 ⟪ spx=px , fc ⟫ = ⊥-elim ( ne spx=px )
1164 zc33 : supf0 z ≡ & (SUP.sup zc32) 1246 zc33 : supf0 z ≡ & (SUP.sup zc32)
1165 zc33 = ? -- trans (sym (supfx (o≤-refl0 (sym zc30)))) ( ZChain.supf-is-minsup zc o≤-refl ) 1247 zc33 = ? -- trans (sym (supfx (o≤-refl0 (sym zc30)))) ( ZChain.supf-is-minsup zc o≤-refl )
1166 zc36 : ¬ (supf0 px ≡ px) → STMP z≤x 1248 zc36 : ¬ (supf0 px ≡ px) → STMP z≤x
1167 zc36 ne = ? -- record { tsup = record { sup = SUP.sup zc32 ; as = SUP.as zc32 ; x≤sup = zc34 ne } ; tsup=sup = zc33 } 1249 zc36 ne = ? -- record { tsup = record { sup = SUP.sup zc32 ; as = SUP.as zc32 ; x≤sup = zc34 ne } ; tsup=sup = zc33 }
1168 zc35 : STMP z≤x 1250 zc35 : STMP z≤x
1169 zc35 with trio< (supf0 px) px 1251 zc35 with trio< (supf0 px) px
1170 ... | tri< a ¬b ¬c = zc36 ¬b 1252 ... | tri< a ¬b ¬c = zc36 ¬b
1171 ... | tri> ¬a ¬b c = zc36 ¬b 1253 ... | tri> ¬a ¬b c = zc36 ¬b
1172 ... | tri≈ ¬a b ¬c = record { tsup = ? ; tsup=sup = ? } where 1254 ... | tri≈ ¬a b ¬c = record { tsup = ? ; tsup=sup = ? } where
1173 zc37 : MinSUP A (UnionCF A f mf ay supf0 z) 1255 zc37 : MinSUP A (UnionCF A f mf ay supf0 z)
1174 zc37 = record { sup = ? ; asm = ? ; x≤sup = ? } 1256 zc37 = record { sup = ? ; asm = ? ; x≤sup = ? }
1175 sup=u : {b : Ordinal} (ab : odef A b) → 1257 sup=u : {b : Ordinal} (ab : odef A b) →
1176 b o≤ x → IsMinSUP A (UnionCF A f mf ay supf0 b) supf0 ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf0 b) f b ) → supf0 b ≡ b 1258 b o≤ x → IsMinSUP A (UnionCF A f mf ay supf0 b) supf0 ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf0 b) f b ) → supf0 b ≡ b
1177 sup=u {b} ab b≤x is-sup with trio< b px 1259 sup=u {b} ab b≤x is-sup with trio< b px
1178 ... | tri< a ¬b ¬c = ZChain.sup=u zc ab (o<→≤ a) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫ 1260 ... | tri< a ¬b ¬c = ZChain.sup=u zc ab (o<→≤ a) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫
1179 ... | tri≈ ¬a b ¬c = ZChain.sup=u zc ab (o≤-refl0 b) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫ 1261 ... | tri≈ ¬a b ¬c = ZChain.sup=u zc ab (o≤-refl0 b) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫
1180 ... | tri> ¬a ¬b px<b = zc31 ? where 1262 ... | tri> ¬a ¬b px<b = zc31 ? where
1181 zc30 : x ≡ b 1263 zc30 : x ≡ b
1182 zc30 with osuc-≡< b≤x 1264 zc30 with osuc-≡< b≤x
1183 ... | case1 eq = sym (eq) 1265 ... | case1 eq = sym (eq)
1184 ... | case2 b<x = ⊥-elim (¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ ) 1266 ... | case2 b<x = ⊥-elim (¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ )
1185 zcsup : xSUP (UnionCF A f mf ay supf0 px) supf0 x 1267 zcsup : xSUP (UnionCF A f mf ay supf0 px) supf0 x
1186 zcsup with zc30 1268 zcsup with zc30
1187 ... | refl = record { ax = ab ; is-sup = record { x≤sup = λ {w} lt → 1269 ... | refl = record { ax = ab ; is-sup = record { x≤sup = λ {w} lt →
1188 IsMinSUP.x≤sup (proj1 is-sup) ? ; minsup = ? } } 1270 IsMinSUP.x≤sup (proj1 is-sup) ? ; minsup = ? } }
1189 zc31 : ( (¬ xSUP (UnionCF A f mf ay supf0 px) supf0 x ) ∨ HasPrev A (UnionCF A f mf ay supf0 px) f x ) → supf0 b ≡ b 1271 zc31 : ( (¬ xSUP (UnionCF A f mf ay supf0 px) supf0 x ) ∨ HasPrev A (UnionCF A f mf ay supf0 px) f x ) → supf0 b ≡ b
1190 zc31 (case1 ¬sp=x) with zc30 1272 zc31 (case1 ¬sp=x) with zc30
1191 ... | refl = ⊥-elim (¬sp=x zcsup ) 1273 ... | refl = ⊥-elim (¬sp=x zcsup )
1192 zc31 (case2 hasPrev ) with zc30 1274 zc31 (case2 hasPrev ) with zc30
1193 ... | refl = ⊥-elim ( proj2 is-sup record { ax = HasPrev.ax hasPrev ; y = HasPrev.y hasPrev 1275 ... | refl = ⊥-elim ( proj2 is-sup record { ax = HasPrev.ax hasPrev ; y = HasPrev.y hasPrev
1194 ; ay = ? ; x=fy = HasPrev.x=fy hasPrev } ) 1276 ; ay = ? ; x=fy = HasPrev.x=fy hasPrev } )
1195 1277
1196 ... | no lim = zc5 where 1278 ... | no lim = zc5 where
1197 1279
1198 pzc : (z : Ordinal) → z o< x → ZChain A f mf ay z 1280 pzc : (z : Ordinal) → z o< x → ZChain A f mf ay z
1199 pzc z z<x = prev z z<x 1281 pzc z z<x = prev z z<x
1208 1290
1209 pchain : HOD 1291 pchain : HOD
1210 pchain = UnionCF A f mf ay supf0 x 1292 pchain = UnionCF A f mf ay supf0 x
1211 1293
1212 ptotal0 : IsTotalOrderSet pchain 1294 ptotal0 : IsTotalOrderSet pchain
1213 ptotal0 {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where 1295 ptotal0 {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
1214 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) 1296 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
1215 uz01 = chain-total A f mf ay supf0 ( (proj2 ca)) ( (proj2 cb)) 1297 uz01 = chain-total A f mf ay supf0 ( (proj2 ca)) ( (proj2 cb))
1216 1298
1217 usup : MinSUP A pchain 1299 usup : MinSUP A pchain
1218 usup = minsupP pchain (λ lt → proj1 lt) ptotal0 1300 usup = minsupP pchain (λ lt → proj1 lt) ptotal0
1219 spu = MinSUP.sup usup 1301 spu = MinSUP.sup usup
1220 1302
1221 supf1 : Ordinal → Ordinal 1303 supf1 : Ordinal → Ordinal
1227 pchain1 : HOD 1309 pchain1 : HOD
1228 pchain1 = UnionCF A f mf ay supf1 x 1310 pchain1 = UnionCF A f mf ay supf1 x
1229 1311
1230 is-max-hp : (supf : Ordinal → Ordinal) (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a → 1312 is-max-hp : (supf : Ordinal → Ordinal) (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a →
1231 b o< x → (ab : odef A b) → 1313 b o< x → (ab : odef A b) →
1232 HasPrev A (UnionCF A f mf ay supf x) f b → 1314 HasPrev A (UnionCF A f mf ay supf x) f b →
1233 * a < * b → odef (UnionCF A f mf ay supf x) b 1315 * a < * b → odef (UnionCF A f mf ay supf x) b
1234 is-max-hp supf x {a} {b} ua b<x ab has-prev a<b with HasPrev.ay has-prev 1316 is-max-hp supf x {a} {b} ua b<x ab has-prev a<b with HasPrev.ay has-prev
1235 ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ? -- ⟪ ab , 1317 ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫
1318 ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ? -- ⟪ ab ,
1236 -- subst (λ k → UChain A f mf ay supf x k ) 1319 -- subst (λ k → UChain A f mf ay supf x k )
1237 -- (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc)) ⟫ 1320 -- (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc)) ⟫
1238 1321
1239 zc70 : HasPrev A pchain f x → ¬ xSUP pchain f x 1322 zc70 : HasPrev A pchain f x → ¬ xSUP pchain f x
1240 zc70 pr xsup = ? 1323 zc70 pr xsup = ?
1241 1324
1242 no-extension : ¬ ( xSUP (UnionCF A f mf ay supf0 x) supf0 x ) → ZChain A f mf ay x 1325 no-extension : ¬ ( xSUP (UnionCF A f mf ay supf0 x) supf0 x ) → ZChain A f mf ay x
1243 no-extension ¬sp=x = ? where -- record { supf = supf1 ; sup=u = sup=u 1326 no-extension ¬sp=x = ? where -- record { supf = supf1 ; sup=u = sup=u
1244 -- ; sup = sup ; supf-is-sup = sis ; supf-mono = {!!} ; asupf = ? } where 1327 -- ; sup = sup ; supf-is-sup = sis ; supf-mono = {!!} ; asupf = ? } where
1245 supfu : {u : Ordinal } → ( a : u o< x ) → (z : Ordinal) → Ordinal 1328 supfu : {u : Ordinal } → ( a : u o< x ) → (z : Ordinal) → Ordinal
1246 supfu {u} a z = ZChain.supf (pzc (osuc u) (ob<x lim a)) z 1329 supfu {u} a z = ZChain.supf (pzc (osuc u) (ob<x lim a)) z
1247 pchain0=1 : pchain ≡ pchain1 1330 pchain0=1 : pchain ≡ pchain1
1248 pchain0=1 = ==→o≡ record { eq→ = zc10 ; eq← = zc11 } where 1331 pchain0=1 = ==→o≡ record { eq→ = zc10 ; eq← = zc11 } where
1249 zc10 : {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchain1) z 1332 zc10 : {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchain1) z
1333 zc10 {z} ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1250 zc10 {z} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ = zc12 fc where 1334 zc10 {z} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ = zc12 fc where
1251 zc12 : {z : Ordinal} → FClosure A f (supf0 u) z → odef pchain1 z 1335 zc12 : {z : Ordinal} → FClosure A f (supf0 u) z → odef pchain1 z
1252 zc12 (fsuc x fc) with zc12 fc 1336 zc12 (fsuc x fc) with zc12 fc
1253 ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫ 1337 ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-init (fsuc _ fc₁) ⟫
1254 zc12 (init asu su=z ) = ⟪ subst (λ k → odef A k) su=z asu , ch-is-sup u ? ? (init ? ? ) ⟫ 1338 ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫
1339 zc12 (init asu su=z ) = ⟪ subst (λ k → odef A k) su=z asu , ch-is-sup u ? ? (init ? ? ) ⟫
1255 zc11 : {z : Ordinal} → OD.def (od pchain1) z → OD.def (od pchain) z 1340 zc11 : {z : Ordinal} → OD.def (od pchain1) z → OD.def (od pchain) z
1341 zc11 {z} ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1256 zc11 {z} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ = zc13 fc where 1342 zc11 {z} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ = zc13 fc where
1257 zc13 : {z : Ordinal} → FClosure A f (supf1 u) z → odef pchain z 1343 zc13 : {z : Ordinal} → FClosure A f (supf1 u) z → odef pchain z
1258 zc13 (fsuc x fc) with zc13 fc 1344 zc13 (fsuc x fc) with zc13 fc
1259 ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫ 1345 ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-init (fsuc _ fc₁) ⟫
1346 ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫
1260 zc13 (init asu su=z ) with trio< u x 1347 zc13 (init asu su=z ) with trio< u x
1261 ... | tri< a ¬b ¬c = ⟪ ? , ch-is-sup u ? ? (init ? ? ) ⟫ 1348 ... | tri< a ¬b ¬c = ⟪ ? , ch-is-sup u ? ? (init ? ? ) ⟫
1262 ... | tri≈ ¬a b ¬c = ? 1349 ... | tri≈ ¬a b ¬c = ?
1263 ... | tri> ¬a ¬b c = ? -- ⊥-elim ( o≤> u<x c ) 1350 ... | tri> ¬a ¬b c = ? -- ⊥-elim ( o≤> u<x c )
1264 sup : {z : Ordinal} → z o≤ x → SUP A (UnionCF A f mf ay supf1 z) 1351 sup : {z : Ordinal} → z o≤ x → SUP A (UnionCF A f mf ay supf1 z)
1265 sup {z} z≤x with trio< z x 1352 sup {z} z≤x with trio< z x
1266 ... | tri< a ¬b ¬c = SUP⊆ ? (ZChain.sup (pzc (osuc z) {!!}) {!!} ) 1353 ... | tri< a ¬b ¬c = SUP⊆ ? (ZChain.sup (pzc (osuc z) {!!}) {!!} )
1267 ... | tri≈ ¬a b ¬c = {!!} 1354 ... | tri≈ ¬a b ¬c = {!!}
1268 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x )) 1355 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x ))
1269 sis : {z : Ordinal} (x≤z : z o≤ x) → supf1 z ≡ & (SUP.sup (sup {!!})) 1356 sis : {z : Ordinal} (x≤z : z o≤ x) → supf1 z ≡ & (SUP.sup (sup {!!}))
1270 sis {z} z≤x with trio< z x 1357 sis {z} z≤x with trio< z x
1271 ... | tri< a ¬b ¬c = {!!} where 1358 ... | tri< a ¬b ¬c = {!!} where
1272 zc8 = ZChain.supf-is-minsup (pzc z a) {!!} 1359 zc8 = ZChain.supf-is-minsup (pzc z a) {!!}
1273 ... | tri≈ ¬a b ¬c = {!!} 1360 ... | tri≈ ¬a b ¬c = {!!}
1274 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x )) 1361 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x ))
1275 sup=u : {b : Ordinal} (ab : odef A b) → b o≤ x → IsMinSUP A (UnionCF A f mf ay supf1 b) f ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf1 b) f b ) → supf1 b ≡ b 1362 sup=u : {b : Ordinal} (ab : odef A b) → b o≤ x → IsMinSUP A (UnionCF A f mf ay supf1 b) f ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf1 b) f b ) → supf1 b ≡ b
1276 sup=u {z} ab z≤x is-sup with trio< z x 1363 sup=u {z} ab z≤x is-sup with trio< z x
1277 ... | tri< a ¬b ¬c = ? -- ZChain.sup=u (pzc (osuc b) (ob<x lim a)) ab {!!} record { x≤sup = {!!} } 1364 ... | tri< a ¬b ¬c = ? -- ZChain.sup=u (pzc (osuc b) (ob<x lim a)) ab {!!} record { x≤sup = {!!} }
1278 ... | tri≈ ¬a b ¬c = {!!} -- ZChain.sup=u (pzc (osuc ?) ?) ab {!!} record { x≤sup = {!!} } 1365 ... | tri≈ ¬a b ¬c = {!!} -- ZChain.sup=u (pzc (osuc ?) ?) ab {!!} record { x≤sup = {!!} }
1279 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x )) 1366 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x ))
1280 1367
1281 zc5 : ZChain A f mf ay x 1368 zc5 : ZChain A f mf ay x
1282 zc5 with ODC.∋-p O A (* x) 1369 zc5 with ODC.∋-p O A (* x)
1283 ... | no noax = no-extension {!!} -- ¬ A ∋ p, just skip 1370 ... | no noax = no-extension {!!} -- ¬ A ∋ p, just skip
1284 ... | yes ax with ODC.p∨¬p O ( HasPrev A pchain f x ) 1371 ... | yes ax with ODC.p∨¬p O ( HasPrev A pchain f x )
1285 -- we have to check adding x preserve is-max ZChain A y f mf x 1372 -- we have to check adding x preserve is-max ZChain A y f mf x
1286 ... | case1 pr = no-extension {!!} 1373 ... | case1 pr = no-extension {!!}
1287 ... | case2 ¬fy<x with ODC.p∨¬p O (IsMinSUP A pchain f ax ) 1374 ... | case2 ¬fy<x with ODC.p∨¬p O (IsMinSUP A pchain f ax )
1288 ... | case1 is-sup = ? -- record { supf = supf1 ; sup=u = {!!} 1375 ... | case1 is-sup = ? -- record { supf = supf1 ; sup=u = {!!}
1289 -- ; sup = {!!} ; supf-is-sup = {!!} ; supf-mono = {!!}; asupf = {!!} } -- where -- x is a sup of (zc ?) 1376 -- ; sup = {!!} ; supf-is-sup = {!!} ; supf-mono = {!!}; asupf = {!!} } -- where -- x is a sup of (zc ?)
1290 ... | case2 ¬x=sup = no-extension {!!} -- x is not f y' nor sup of former ZChain from y -- no extention 1377 ... | case2 ¬x=sup = no-extension {!!} -- x is not f y' nor sup of former ZChain from y -- no extention
1291 1378
1292 --- 1379 ---
1293 --- the maximum chain has fix point of any ≤-monotonic function 1380 --- the maximum chain has fix point of any ≤-monotonic function
1294 --- 1381 ---
1295 1382
1296 SZ : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → (x : Ordinal) → ZChain A f mf ay x 1383 SZ : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → (x : Ordinal) → ZChain A f mf ay x
1297 SZ f mf {y} ay x = TransFinite {λ z → ZChain A f mf ay z } (λ x → ind f mf ay x ) x 1384 SZ f mf {y} ay x = TransFinite {λ z → ZChain A f mf ay z } (λ x → ind f mf ay x ) x
1298 1385
1299 msp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {x y : Ordinal} (ay : odef A y) 1386 msp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {x y : Ordinal} (ay : odef A y)
1300 → (zc : ZChain A f mf ay x ) 1387 → (zc : ZChain A f mf ay x )
1301 → MinSUP A (UnionCF A f mf ay (ZChain.supf zc) x) 1388 → MinSUP A (UnionCF A f mf ay (ZChain.supf zc) x)
1302 msp0 f mf {x} ay zc = minsupP (UnionCF A f mf ay (ZChain.supf zc) x) (ZChain.chain⊆A zc) (ZChain.f-total zc) 1389 msp0 f mf {x} ay zc = minsupP (UnionCF A f mf ay (ZChain.supf zc) x) (ZChain.chain⊆A zc) (ZChain.f-total zc)
1303 1390
1304 fixpoint : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (zc : ZChain A f mf as0 (& A) ) 1391 fixpoint : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (zc : ZChain A f mf as0 (& A) )
1305 → (sp1 : MinSUP A (ZChain.chain zc)) 1392 → (sp1 : MinSUP A (ZChain.chain zc))
1306 → (ssp<as : ZChain.supf zc (MinSUP.sup sp1) o< ZChain.supf zc (& A)) 1393 → (ssp<as : ZChain.supf zc (MinSUP.sup sp1) o< ZChain.supf zc (& A))
1307 → f (MinSUP.sup sp1) ≡ MinSUP.sup sp1 1394 → f (MinSUP.sup sp1) ≡ MinSUP.sup sp1
1308 fixpoint f mf zc sp1 ssp<as = z14 where 1395 fixpoint f mf zc sp1 ssp<as = z14 where
1309 chain = ZChain.chain zc 1396 chain = ZChain.chain zc
1310 supf = ZChain.supf zc 1397 supf = ZChain.supf zc
1311 sp : Ordinal 1398 sp : Ordinal
1312 sp = MinSUP.sup sp1 1399 sp = MinSUP.sup sp1
1313 asp : odef A sp 1400 asp : odef A sp
1314 asp = MinSUP.asm sp1 1401 asp = MinSUP.asm sp1
1315 z10 : {a b : Ordinal } → (ca : odef chain a ) → supf b o< supf (& A) → (ab : odef A b ) 1402 z10 : {a b : Ordinal } → (ca : odef chain a ) → supf b o< supf (& A) → (ab : odef A b )
1316 → HasPrev A chain f b ∨ IsSUP A (UnionCF A f mf as0 (ZChain.supf zc) b) ab 1403 → HasPrev A chain f b ∨ IsSUP A (UnionCF A f mf as0 (ZChain.supf zc) b) ab
1317 → * a < * b → odef chain b 1404 → * a < * b → odef chain b
1318 z10 = ZChain1.is-max (SZ1 f mf as0 zc (& A) ) 1405 z10 = ZChain1.is-max (SZ1 f mf as0 zc (& A) )
1319 z22 : sp o< & A 1406 z22 : sp o< & A
1320 z22 = z09 asp 1407 z22 = z09 asp
1323 ... | yes eq = subst (λ k → odef chain k) eq ( ZChain.chain∋init zc ) 1410 ... | yes eq = subst (λ k → odef chain k) eq ( ZChain.chain∋init zc )
1324 ... | no ne = ZChain1.is-max (SZ1 f mf as0 zc (& A)) {& s} {sp} ( ZChain.chain∋init zc ) ssp<as asp (case2 z19 ) z13 where 1411 ... | no ne = ZChain1.is-max (SZ1 f mf as0 zc (& A)) {& s} {sp} ( ZChain.chain∋init zc ) ssp<as asp (case2 z19 ) z13 where
1325 z13 : * (& s) < * sp 1412 z13 : * (& s) < * sp
1326 z13 with MinSUP.x≤sup sp1 ( ZChain.chain∋init zc ) 1413 z13 with MinSUP.x≤sup sp1 ( ZChain.chain∋init zc )
1327 ... | case1 eq = ⊥-elim ( ne eq ) 1414 ... | case1 eq = ⊥-elim ( ne eq )
1328 ... | case2 lt = lt 1415 ... | case2 lt = lt
1329 z19 : IsSUP A (UnionCF A f mf as0 (ZChain.supf zc) sp) asp 1416 z19 : IsSUP A (UnionCF A f mf as0 (ZChain.supf zc) sp) asp
1330 z19 = record { x≤sup = z20 } where 1417 z19 = record { x≤sup = z20 } where
1331 z20 : {y : Ordinal} → odef (UnionCF A f mf as0 (ZChain.supf zc) sp) y → (y ≡ sp) ∨ (y << sp) 1418 z20 : {y : Ordinal} → odef (UnionCF A f mf as0 (ZChain.supf zc) sp) y → (y ≡ sp) ∨ (y << sp)
1332 z20 {y} zy with MinSUP.x≤sup sp1 1419 z20 {y} zy with MinSUP.x≤sup sp1
1333 (subst (λ k → odef chain k ) (sym &iso) (chain-mono f mf as0 supf (ZChain.supf-mono zc) (o<→≤ z22) zy )) 1420 (subst (λ k → odef chain k ) (sym &iso) (chain-mono f mf as0 supf (ZChain.supf-mono zc) (o<→≤ z22) zy ))
1334 ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso y=p ) 1421 ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso y=p )
1335 ... | case2 y<p = case2 (subst (λ k → * k < _ ) &iso y<p ) 1422 ... | case2 y<p = case2 (subst (λ k → * k < _ ) &iso y<p )
1336 z14 : f sp ≡ sp 1423 z14 : f sp ≡ sp
1337 z14 with ZChain.f-total zc (subst (λ k → odef chain k) (sym &iso) (ZChain.f-next zc z12 )) (subst (λ k → odef chain k) (sym &iso) z12 ) 1424 z14 with ZChain.f-total zc (subst (λ k → odef chain k) (sym &iso) (ZChain.f-next zc z12 )) (subst (λ k → odef chain k) (sym &iso) z12 )
1338 ... | tri< a ¬b ¬c = ⊥-elim z16 where 1425 ... | tri< a ¬b ¬c = ⊥-elim z16 where
1339 z16 : ⊥ 1426 z16 : ⊥
1340 z16 with proj1 (mf (( MinSUP.sup sp1)) ( MinSUP.asm sp1 )) 1427 z16 with proj1 (mf (( MinSUP.sup sp1)) ( MinSUP.asm sp1 ))
1341 ... | case1 eq = ⊥-elim (¬b (sym eq) ) 1428 ... | case1 eq = ⊥-elim (¬b (sym eq) )
1342 ... | case2 lt = ⊥-elim (¬c lt ) 1429 ... | case2 lt = ⊥-elim (¬c lt )
1343 ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) b ) 1430 ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) b )
1344 ... | tri> ¬a ¬b c = ⊥-elim z17 where 1431 ... | tri> ¬a ¬b c = ⊥-elim z17 where
1345 z15 : (f sp ≡ MinSUP.sup sp1) ∨ (* (f sp) < * (MinSUP.sup sp1) ) 1432 z15 : (f sp ≡ MinSUP.sup sp1) ∨ (* (f sp) < * (MinSUP.sup sp1) )
1346 z15 = MinSUP.x≤sup sp1 (ZChain.f-next zc z12 ) 1433 z15 = MinSUP.x≤sup sp1 (ZChain.f-next zc z12 )
1347 z17 : ⊥ 1434 z17 : ⊥
1348 z17 with z15 1435 z17 with z15
1365 -- ZChain forces fix point on any ≤-monotonic function (fixpoint) 1452 -- ZChain forces fix point on any ≤-monotonic function (fixpoint)
1366 -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain 1453 -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain
1367 -- 1454 --
1368 1455
1369 z04 : (nmx : ¬ Maximal A ) → (zc : ZChain A (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A)) → ⊥ 1456 z04 : (nmx : ¬ Maximal A ) → (zc : ZChain A (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A)) → ⊥
1370 z04 nmx zc = <-irr0 {* (cf nmx c)} {* c} 1457 z04 nmx zc = <-irr0 {* (cf nmx c)} {* c}
1371 (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (MinSUP.asm msp1 )))) 1458 (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (MinSUP.asm msp1 ))))
1372 (subst (λ k → odef A k) (sym &iso) (MinSUP.asm msp1) ) 1459 (subst (λ k → odef A k) (sym &iso) (MinSUP.asm msp1) )
1373 (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) zc msp1 ss<sa ))) -- x ≡ f x ̄ 1460 (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) zc msp1 ss<sa ))) -- x ≡ f x ̄
1374 (proj1 (cf-is-<-monotonic nmx c (MinSUP.asm msp1 ))) where -- x < f x 1461 (proj1 (cf-is-<-monotonic nmx c (MinSUP.asm msp1 ))) where -- x < f x
1375 1462
1376 supf = ZChain.supf zc 1463 supf = ZChain.supf zc
1377 msp1 : MinSUP A (ZChain.chain zc) 1464 msp1 : MinSUP A (ZChain.chain zc)
1378 msp1 = msp0 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc 1465 msp1 = msp0 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc
1379 c : Ordinal 1466 c : Ordinal
1380 c = MinSUP.sup msp1 1467 c = MinSUP.sup msp1
1381 mc = c 1468 mc = c
1382 mc<A : mc o< & A 1469 mc<A : mc o< & A
1383 mc<A = ∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫ 1470 mc<A = ∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫
1384 c=mc : c ≡ mc 1471 c=mc : c ≡ mc
1385 c=mc = refl 1472 c=mc = refl
1386 z20 : mc << cf nmx mc 1473 z20 : mc << cf nmx mc
1387 z20 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1) ) 1474 z20 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1) )
1388 asc : odef A (supf mc) 1475 asc : odef A (supf mc)
1389 asc = ZChain.asupf zc 1476 asc = ZChain.asupf zc
1390 spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc ) 1477 spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc )
1391 spd = ysup (cf nmx) (cf-is-≤-monotonic nmx) asc 1478 spd = ysup (cf nmx) (cf-is-≤-monotonic nmx) asc
1392 d = MinSUP.sup spd 1479 d = MinSUP.sup spd
1393 d<A : d o< & A 1480 d<A : d o< & A
1394 d<A = ∈∧P→o< ⟪ MinSUP.asm spd , lift true ⟫ 1481 d<A = ∈∧P→o< ⟪ MinSUP.asm spd , lift true ⟫
1395 msup : MinSUP A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 supf d) 1482 msup : MinSUP A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 supf d)
1396 msup = ZChain.minsup zc (o<→≤ d<A) 1483 msup = ZChain.minsup zc (o<→≤ d<A)
1397 sd=ms : supf d ≡ MinSUP.sup ( ZChain.minsup zc (o<→≤ d<A) ) 1484 sd=ms : supf d ≡ MinSUP.sup ( ZChain.minsup zc (o<→≤ d<A) )
1398 sd=ms = ZChain.supf-is-minsup zc (o<→≤ d<A) 1485 sd=ms = ZChain.supf-is-minsup zc (o<→≤ d<A)
1399 1486
1400 sc<<d : {mc : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc )) 1487 sc<<d : {mc : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc ))
1401 → supf mc << MinSUP.sup spd 1488 → supf mc << MinSUP.sup spd
1402 sc<<d {mc} asc spd = z25 where 1489 sc<<d {mc} asc spd = z25 where
1403 d1 : Ordinal 1490 d1 : Ordinal
1404 d1 = MinSUP.sup spd -- supf d1 ≡ d 1491 d1 = MinSUP.sup spd -- supf d1 ≡ d
1405 z24 : (supf mc ≡ d1) ∨ ( supf mc << d1 ) 1492 z24 : (supf mc ≡ d1) ∨ ( supf mc << d1 )
1415 -- supf mc ≡ d1 1502 -- supf mc ≡ d1
1416 z32 : ((cf nmx (supf mc)) ≡ d1) ∨ ( (cf nmx (supf mc)) << d1 ) 1503 z32 : ((cf nmx (supf mc)) ≡ d1) ∨ ( (cf nmx (supf mc)) << d1 )
1417 z32 = MinSUP.x≤sup spd (fsuc _ (init asc refl)) 1504 z32 = MinSUP.x≤sup spd (fsuc _ (init asc refl))
1418 z29 : (* (cf nmx d1) ≡ * d1) ∨ (* (cf nmx d1) < * d1) 1505 z29 : (* (cf nmx d1) ≡ * d1) ∨ (* (cf nmx d1) < * d1)
1419 z29 with z32 1506 z29 with z32
1420 ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1) ) 1507 ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1) )
1421 ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt) 1508 ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt)
1422 1509
1423 fsc<<d : {mc z : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc )) 1510 fsc<<d : {mc z : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc ))
1424 → (fc : FClosure A (cf nmx) (supf mc) z) → z << MinSUP.sup spd 1511 → (fc : FClosure A (cf nmx) (supf mc) z) → z << MinSUP.sup spd
1425 fsc<<d {mc} {z} asc spd fc = z25 where 1512 fsc<<d {mc} {z} asc spd fc = z25 where
1426 d1 : Ordinal 1513 d1 : Ordinal
1427 d1 = MinSUP.sup spd -- supf d1 ≡ d 1514 d1 = MinSUP.sup spd -- supf d1 ≡ d
1428 z24 : (z ≡ d1) ∨ ( z << d1 ) 1515 z24 : (z ≡ d1) ∨ ( z << d1 )
1438 -- supf mc ≡ d1 1525 -- supf mc ≡ d1
1439 z32 : ((cf nmx z) ≡ d1) ∨ ( (cf nmx z) << d1 ) 1526 z32 : ((cf nmx z) ≡ d1) ∨ ( (cf nmx z) << d1 )
1440 z32 = MinSUP.x≤sup spd (fsuc _ fc) 1527 z32 = MinSUP.x≤sup spd (fsuc _ fc)
1441 z29 : (* (cf nmx d1) ≡ * d1) ∨ (* (cf nmx d1) < * d1) 1528 z29 : (* (cf nmx d1) ≡ * d1) ∨ (* (cf nmx d1) < * d1)
1442 z29 with z32 1529 z29 with z32
1443 ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1) ) 1530 ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1) )
1444 ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt) 1531 ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt)
1445 1532
1446 smc<<d : supf mc << d 1533 smc<<d : supf mc << d
1447 smc<<d = sc<<d asc spd 1534 smc<<d = sc<<d asc spd
1448 1535
1449 sz<<c : {z : Ordinal } → z o< & A → supf z <= mc 1536 sz<<c : {z : Ordinal } → z o< & A → supf z <= mc
1450 sz<<c z<A = MinSUP.x≤sup msp1 (ZChain.csupf zc (z09 (ZChain.asupf zc) )) 1537 sz<<c z<A = MinSUP.x≤sup msp1 (ZChain.csupf zc (z09 (ZChain.asupf zc) ))
1451 1538
1452 sc=c : supf mc ≡ mc 1539 sc=c : supf mc ≡ mc
1453 sc=c = ZChain.sup=u zc (MinSUP.asm msp1) (o<→≤ (∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫ )) ⟪ is-sup , not-hasprev ⟫ where 1540 sc=c = ZChain.sup=u zc (MinSUP.asm msp1) (o<→≤ (∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫ )) ⟪ is-sup , not-hasprev ⟫ where
1454 not-hasprev : ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) mc) (cf nmx) mc 1541 not-hasprev : ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) mc) (cf nmx) mc
1542 not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-init fc ⟫ ; x=fy = x=fy } = ⊥-elim ( <-irr z31 z30 ) where
1543 z30 : * mc < * (cf nmx mc)
1544 z30 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1))
1545 z31 : ( * (cf nmx mc) ≡ * mc ) ∨ ( * (cf nmx mc) < * mc )
1546 z31 = <=to≤ ( MinSUP.x≤sup msp1 (subst (λ k → odef (ZChain.chain zc) (cf nmx k)) (sym x=fy)
1547 ⟪ proj2 (cf-is-≤-monotonic nmx _ (proj2 (cf-is-≤-monotonic nmx _ ua1 ) )) , ch-init (fsuc _ (fsuc _ fc)) ⟫ ))
1455 not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z48 z32 ) where 1548 not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z48 z32 ) where
1456 z30 : * mc < * (cf nmx mc) 1549 z30 : * mc < * (cf nmx mc)
1457 z30 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1)) 1550 z30 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1))
1458 z31 : ( supf mc ≡ mc ) ∨ ( * (supf mc) < * mc ) 1551 z31 : ( supf mc ≡ mc ) ∨ ( * (supf mc) < * mc )
1459 z31 = MinSUP.x≤sup msp1 (ZChain.csupf zc (z09 (ZChain.asupf zc) )) 1552 z31 = MinSUP.x≤sup msp1 (ZChain.csupf zc (z09 (ZChain.asupf zc) ))
1464 is-sup : IsSUP A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) mc) (MinSUP.asm msp1) 1557 is-sup : IsSUP A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) mc) (MinSUP.asm msp1)
1465 is-sup = record { x≤sup = λ zy → MinSUP.x≤sup msp1 (chain-mono (cf nmx) (cf-is-≤-monotonic nmx) as0 supf (ZChain.supf-mono zc) (o<→≤ mc<A) zy ) } 1558 is-sup = record { x≤sup = λ zy → MinSUP.x≤sup msp1 (chain-mono (cf nmx) (cf-is-≤-monotonic nmx) as0 supf (ZChain.supf-mono zc) (o<→≤ mc<A) zy ) }
1466 1559
1467 1560
1468 not-hasprev : ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 supf d) (cf nmx) d 1561 not-hasprev : ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 supf d) (cf nmx) d
1562 not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-init fc ⟫ ; x=fy = x=fy } = ⊥-elim ( <-irr z31 z30 ) where
1563 z30 : * d < * (cf nmx d)
1564 z30 = proj1 (cf-is-<-monotonic nmx d (MinSUP.asm spd))
1565 z32 : ( cf nmx (cf nmx y) ≡ supf mc ) ∨ ( * (cf nmx (cf nmx y)) < * (supf mc) )
1566 z32 = ZChain.fcy<sup zc (o<→≤ mc<A) (fsuc _ (fsuc _ fc))
1567 z31 : ( * (cf nmx d) ≡ * d ) ∨ ( * (cf nmx d) < * d )
1568 z31 = case2 ( subst (λ k → * (cf nmx k) < * d ) (sym x=fy) ( ftrans<=-< z32 ( sc<<d {mc} asc spd ) ))
1469 not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z46 z30 ) where 1569 not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z46 z30 ) where
1470 z45 : (* (cf nmx (cf nmx y)) ≡ * d) ∨ (* (cf nmx (cf nmx y)) < * d) → (* (cf nmx d) ≡ * d) ∨ (* (cf nmx d) < * d) 1570 z45 : (* (cf nmx (cf nmx y)) ≡ * d) ∨ (* (cf nmx (cf nmx y)) < * d) → (* (cf nmx d) ≡ * d) ∨ (* (cf nmx d) < * d)
1471 z45 p = subst (λ k → (* (cf nmx k) ≡ * d) ∨ (* (cf nmx k) < * d)) (sym x=fy) p 1571 z45 p = subst (λ k → (* (cf nmx k) ≡ * d) ∨ (* (cf nmx k) < * d)) (sym x=fy) p
1472 z48 : supf mc << d 1572 z48 : supf mc << d
1473 z48 = sc<<d {mc} asc spd 1573 z48 = sc<<d {mc} asc spd
1474 z53 : supf u o< supf (& A) 1574 z53 : supf u o< supf (& A)
1475 z53 = ordtrans<-≤ u<x (ZChain.supf-mono zc (o<→≤ d<A) ) 1575 z53 = ordtrans<-≤ u<x (ZChain.supf-mono zc (o<→≤ d<A) )
1476 z52 : ( u ≡ mc ) ∨ ( u << mc ) 1576 z52 : ( u ≡ mc ) ∨ ( u << mc )
1477 z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc) 1577 z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc)
1478 , ch-is-sup u z53 is-sup1 (init (ZChain.asupf zc) (ChainP.supu=u is-sup1)) ⟫ 1578 , ch-is-sup u z53 is-sup1 (init (ZChain.asupf zc) (ChainP.supu=u is-sup1)) ⟫
1479 z51 : supf u o≤ supf mc 1579 z51 : supf u o≤ supf mc
1480 z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57 where 1580 z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57 where
1481 z56 : u ≡ mc → supf u ≡ supf mc 1581 z56 : u ≡ mc → supf u ≡ supf mc
1482 z56 eq = cong supf eq 1582 z56 eq = cong supf eq
1483 z57 : u << mc → supf u o≤ supf mc 1583 z57 : u << mc → supf u o≤ supf mc
1484 z57 lt = ZChain.supf-<= zc (case2 z58) where 1584 z57 lt = ZChain.supf-<= zc (case2 z58) where
1485 z58 : supf u << supf mc -- supf u o< supf d -- supf u << supf d 1585 z58 : supf u << supf mc -- supf u o< supf d -- supf u << supf d
1486 z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c) lt 1586 z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c) lt
1487 z49 : supf u o< supf mc → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) ) 1587 z49 : supf u o< supf mc → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) )
1488 z49 su<smc = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A su<smc (fsuc _ ( fsuc _ fc )) 1588 z49 su<smc = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A su<smc (fsuc _ ( fsuc _ fc ))
1489 z50 : (cf nmx (cf nmx y) ≡ supf d) ∨ (* (cf nmx (cf nmx y)) < * (supf d) ) 1589 z50 : (cf nmx (cf nmx y) ≡ supf d) ∨ (* (cf nmx (cf nmx y)) < * (supf d) )
1490 z50 = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) d<A u<x (fsuc _ ( fsuc _ fc )) 1590 z50 = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) d<A u<x (fsuc _ ( fsuc _ fc ))
1491 z47 : {mc d1 : Ordinal } {asc : odef A (supf mc)} (spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc )) 1591 z47 : {mc d1 : Ordinal } {asc : odef A (supf mc)} (spd : MinSUP A (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc ))
1492 → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) ) → supf mc << d1 1592 → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) ) → supf mc << d1
1493 → * (cf nmx (cf nmx y)) < * d1 1593 → * (cf nmx (cf nmx y)) < * d1
1494 z47 {mc} {d1} {asc} spd (case1 eq) smc<d = subst (λ k → k < * d1 ) (sym (cong (*) eq)) smc<d 1594 z47 {mc} {d1} {asc} spd (case1 eq) smc<d = subst (λ k → k < * d1 ) (sym (cong (*) eq)) smc<d
1495 z47 {mc} {d1} {asc} spd (case2 lt) smc<d = IsStrictPartialOrder.trans PO lt smc<d 1595 z47 {mc} {d1} {asc} spd (case2 lt) smc<d = IsStrictPartialOrder.trans PO lt smc<d
1496 z30 : * d < * (cf nmx d) 1596 z30 : * d < * (cf nmx d)
1497 z30 = proj1 (cf-is-<-monotonic nmx d (MinSUP.asm spd)) 1597 z30 = proj1 (cf-is-<-monotonic nmx d (MinSUP.asm spd))
1498 z46 : (* (cf nmx d) ≡ * d) ∨ (* (cf nmx d) < * d) 1598 z46 : (* (cf nmx d) ≡ * d) ∨ (* (cf nmx d) < * d)
1499 z46 = or (osuc-≡< z51) z55 z54 where 1599 z46 = or (osuc-≡< z51) z55 z54 where
1509 is-sup = record { x≤sup = z22 } where 1609 is-sup = record { x≤sup = z22 } where
1510 z23 : {z : Ordinal } → odef (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc) z → (z ≡ MinSUP.sup spd) ∨ (z << MinSUP.sup spd) 1610 z23 : {z : Ordinal } → odef (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc) z → (z ≡ MinSUP.sup spd) ∨ (z << MinSUP.sup spd)
1511 z23 lt = MinSUP.x≤sup spd lt 1611 z23 lt = MinSUP.x≤sup spd lt
1512 z22 : {y : Ordinal} → odef (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) d) y → 1612 z22 : {y : Ordinal} → odef (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) d) y →
1513 (y ≡ MinSUP.sup spd) ∨ (y << MinSUP.sup spd) 1613 (y ≡ MinSUP.sup spd) ∨ (y << MinSUP.sup spd)
1614 z22 {a} ⟪ aa , ch-init fc ⟫ = case2 ( ( ftrans<=-< z32 ( sc<<d {mc} asc spd ) )) where
1615 z32 : ( a ≡ supf mc ) ∨ ( * a < * (supf mc) )
1616 z32 = ZChain.fcy<sup zc (o<→≤ mc<A) fc
1514 z22 {a} ⟪ aa , ch-is-sup u u<x is-sup1 fc ⟫ = tri u (supf mc) 1617 z22 {a} ⟪ aa , ch-is-sup u u<x is-sup1 fc ⟫ = tri u (supf mc)
1515 z60 z61 ( λ sc<u → ⊥-elim ( o≤> ( subst (λ k → k o≤ supf mc) (ChainP.supu=u is-sup1) z51) sc<u )) where 1618 z60 z61 ( λ sc<u → ⊥-elim ( o≤> ( subst (λ k → k o≤ supf mc) (ChainP.supu=u is-sup1) z51) sc<u )) where
1516 z53 : supf u o< supf (& A) 1619 z53 : supf u o< supf (& A)
1517 z53 = ordtrans<-≤ u<x (ZChain.supf-mono zc (o<→≤ d<A) ) 1620 z53 = ordtrans<-≤ u<x (ZChain.supf-mono zc (o<→≤ d<A) )
1518 z52 : ( u ≡ mc ) ∨ ( u << mc ) 1621 z52 : ( u ≡ mc ) ∨ ( u << mc )
1519 z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc) 1622 z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc)
1520 , ch-is-sup u z53 is-sup1 (init (ZChain.asupf zc) (ChainP.supu=u is-sup1)) ⟫ 1623 , ch-is-sup u z53 is-sup1 (init (ZChain.asupf zc) (ChainP.supu=u is-sup1)) ⟫
1521 z56 : u ≡ mc → supf u ≡ supf mc 1624 z56 : u ≡ mc → supf u ≡ supf mc
1522 z56 eq = cong supf eq 1625 z56 eq = cong supf eq
1523 z57 : u << mc → supf u o≤ supf mc 1626 z57 : u << mc → supf u o≤ supf mc
1524 z57 lt = ZChain.supf-<= zc (case2 z58) where 1627 z57 lt = ZChain.supf-<= zc (case2 z58) where
1525 z58 : supf u << supf mc -- supf u o< supf d -- supf u << supf d 1628 z58 : supf u << supf mc -- supf u o< supf d -- supf u << supf d
1526 z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c) lt 1629 z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c) lt
1527 z51 : supf u o≤ supf mc 1630 z51 : supf u o≤ supf mc
1528 z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57 1631 z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57
1529 z60 : u o< supf mc → (a ≡ d ) ∨ ( * a < * d ) 1632 z60 : u o< supf mc → (a ≡ d ) ∨ ( * a < * d )
1530 z60 u<smc = case2 ( ftrans<=-< (ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A 1633 z60 u<smc = case2 ( ftrans<=-< (ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A
1531 (subst (λ k → k o< supf mc) (sym (ChainP.supu=u is-sup1)) u<smc) fc ) smc<<d ) 1634 (subst (λ k → k o< supf mc) (sym (ChainP.supu=u is-sup1)) u<smc) fc ) smc<<d )
1532 z61 : u ≡ supf mc → (a ≡ d ) ∨ ( * a < * d ) 1635 z61 : u ≡ supf mc → (a ≡ d ) ∨ ( * a < * d )
1533 z61 u=sc = case2 (fsc<<d {mc} asc spd (subst (λ k → FClosure A (cf nmx) k a) (trans (ChainP.supu=u is-sup1) u=sc) fc ) ) 1636 z61 u=sc = case2 (fsc<<d {mc} asc spd (subst (λ k → FClosure A (cf nmx) k a) (trans (ChainP.supu=u is-sup1) u=sc) fc ) )
1534 -- u<x : ZChain.supf zc u o< ZChain.supf zc d 1637 -- u<x : ZChain.supf zc u o< ZChain.supf zc d
1535 -- supf u o< spuf c → order 1638 -- supf u o< spuf c → order
1543 ... | case2 lt = lt 1646 ... | case2 lt = lt
1544 1647
1545 sms<sa : supf mc o< supf (& A) 1648 sms<sa : supf mc o< supf (& A)
1546 sms<sa with osuc-≡< ( ZChain.supf-mono zc (o<→≤ ( ∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫) )) 1649 sms<sa with osuc-≡< ( ZChain.supf-mono zc (o<→≤ ( ∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫) ))
1547 ... | case2 lt = lt 1650 ... | case2 lt = lt
1548 ... | case1 eq = ⊥-elim ( o<¬≡ eq ( ordtrans<-≤ (sc<sd (subst (λ k → supf mc << k ) (sym sd=d) (sc<<d {mc} asc spd)) ) 1651 ... | case1 eq = ⊥-elim ( o<¬≡ eq ( ordtrans<-≤ (sc<sd (subst (λ k → supf mc << k ) (sym sd=d) (sc<<d {mc} asc spd)) )
1549 ( ZChain.supf-mono zc (o<→≤ d<A )))) 1652 ( ZChain.supf-mono zc (o<→≤ d<A ))))
1550 1653
1551 ss<sa : supf c o< supf (& A) 1654 ss<sa : supf c o< supf (& A)
1552 ss<sa = subst (λ k → supf k o< supf (& A)) (sym c=mc) sms<sa 1655 ss<sa = subst (λ k → supf k o< supf (& A)) (sym c=mc) sms<sa
1553 1656
1554 zorn00 : Maximal A 1657 zorn00 : Maximal A
1555 zorn00 with is-o∅ ( & HasMaximal ) -- we have no Level (suc n) LEM 1658 zorn00 with is-o∅ ( & HasMaximal ) -- we have no Level (suc n) LEM
1556 ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; as = zorn01 ; ¬maximal<x = zorn02 } where 1659 ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; as = zorn01 ; ¬maximal<x = zorn02 } where
1557 -- yes we have the maximal 1660 -- yes we have the maximal
1558 zorn03 : odef HasMaximal ( & ( ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ) ) 1661 zorn03 : odef HasMaximal ( & ( ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ) )
1559 zorn03 = ODC.x∋minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) -- Axiom of choice 1662 zorn03 = ODC.x∋minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) -- Axiom of choice
1560 zorn01 : A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) 1663 zorn01 : A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
1561 zorn01 = proj1 zorn03 1664 zorn01 = proj1 zorn03
1562 zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x) 1665 zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x)
1563 zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x ) 1666 zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x )
1564 ... | yes ¬Maximal = ⊥-elim ( z04 nmx (SZ (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A) )) where 1667 ... | yes ¬Maximal = ⊥-elim ( z04 nmx (SZ (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A) )) where
1565 -- if we have no maximal, make ZChain, which contradict SUP condition 1668 -- if we have no maximal, make ZChain, which contradict SUP condition
1566 nmx : ¬ Maximal A 1669 nmx : ¬ Maximal A
1567 nmx mx = ∅< {HasMaximal} zc5 ( ≡o∅→=od∅ ¬Maximal ) where 1670 nmx mx = ∅< {HasMaximal} zc5 ( ≡o∅→=od∅ ¬Maximal ) where
1568 zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) → odef A y → ¬ (* (& (Maximal.maximal mx)) < * y)) 1671 zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) → odef A y → ¬ (* (& (Maximal.maximal mx)) < * y))
1569 zc5 = ⟪ Maximal.as mx , (λ y ay mx<y → Maximal.¬maximal<x mx (subst (λ k → odef A k ) (sym &iso) ay) (subst (λ k → k < * y) *iso mx<y) ) ⟫ 1672 zc5 = ⟪ Maximal.as mx , (λ y ay mx<y → Maximal.¬maximal<x mx (subst (λ k → odef A k ) (sym &iso) ay) (subst (λ k → k < * y) *iso mx<y) ) ⟫
1570 1673
1571 -- usage (see filter.agda ) 1674 -- usage (see filter.agda )
1572 -- 1675 --
1573 -- _⊆'_ : ( A B : HOD ) → Set n 1676 -- _⊆'_ : ( A B : HOD ) → Set n
1574 -- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x 1677 -- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x
1575 1678
1576 -- MaximumSubset : {L P : HOD} 1679 -- MaximumSubset : {L P : HOD}
1577 -- → o∅ o< & L → o∅ o< & P → P ⊆ L 1680 -- → o∅ o< & L → o∅ o< & P → P ⊆ L
1578 -- → IsPartialOrderSet P _⊆'_ 1681 -- → IsPartialOrderSet P _⊆'_
1579 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ ) 1682 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ )
1580 -- → Maximal P (_⊆'_) 1683 -- → Maximal P (_⊆'_)
1581 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆'_} 0<P PO SP 1684 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆'_} 0<P PO SP