annotate src/zorn.agda @ 625:886e1f82cca0

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 20 Jun 2022 08:43:23 +0900
parents d0938f220648
children 35d8aca1a2b7
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
478
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 477
diff changeset
1 {-# OPTIONS --allow-unsolved-metas #-}
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
2 open import Level hiding ( suc ; zero )
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
3 open import Ordinals
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
4 open import Relation.Binary
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
5 open import Relation.Binary.Core
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
6 open import Relation.Binary.PropositionalEquality
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
7 import OD
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
8 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
9
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
10 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
11 -- Zorn-lemma : { A : HOD }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
12 -- → o∅ o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
13 -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
14 -- → Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
15 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
16
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
17 open import zf
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
18 open import logic
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
19 -- open import partfunc {n} O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
20
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
21 open import Relation.Nullary
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
22 open import Data.Empty
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
23 import BAlgbra
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
24
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
25 open import Data.Nat hiding ( _<_ ; _≤_ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
26 open import Data.Nat.Properties
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
27 open import nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
28
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
29
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
30 open inOrdinal O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
31 open OD O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
32 open OD.OD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 open ODAxiom odAxiom
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
34 import OrdUtil
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
35 import ODUtil
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
36 open Ordinals.Ordinals O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
37 open Ordinals.IsOrdinals isOrdinal
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
38 open Ordinals.IsNext isNext
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
39 open OrdUtil O
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
40 open ODUtil O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
41
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
42
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
43 import ODC
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
44
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
45
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
46 open _∧_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
47 open _∨_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
48 open Bool
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
49
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
50
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
51 open HOD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
52
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
53 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
54 -- Partial Order on HOD ( possibly limited in A )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
55 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
56
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
57 _<<_ : (x y : Ordinal ) → Set n -- Set n order
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
58 x << y = * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
59
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
60 POO : IsStrictPartialOrder _≡_ _<<_
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
61 POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
62 ; trans = IsStrictPartialOrder.trans PO
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
63 ; irrefl = λ x=y x<y → IsStrictPartialOrder.irrefl PO (cong (*) x=y) x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
64 ; <-resp-≈ = record { fst = λ {x} {y} {y1} y=y1 xy1 → subst (λ k → x << k ) y=y1 xy1 ; snd = λ {x} {x1} {y} x=x1 x1y → subst (λ k → k << x ) x=x1 x1y } }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
65
528
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
66 _≤_ : (x y : HOD) → Set (Level.suc n)
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
67 x ≤ y = ( x ≡ y ) ∨ ( x < y )
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
68
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
69 ≤-ftrans : {x y z : HOD} → x ≤ y → y ≤ z → x ≤ z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
70 ≤-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
71 ≤-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
72 ≤-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
73 ≤-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
74
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
75 <-irr : {a b : HOD} → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
76 <-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (sym a=b) b<a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
77 <-irr {a} {b} (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
78 (IsStrictPartialOrder.trans PO b<a a<b)
490
00c71d1dc316 IsPartialOrder
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 489
diff changeset
79
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
80 ptrans = IsStrictPartialOrder.trans PO
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
81
492
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
82 open _==_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
83 open _⊆_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
84
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
85 --
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
86 -- Closure of ≤-monotonic function f has total order
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
87 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
88
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
89 ≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set (Level.suc n)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
90 ≤-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x ≤ * (f x) ) ∧ odef A (f x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
91
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
92 data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
93 init : odef A s → FClosure A f s s
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
94 fsuc : (x : Ordinal) ( p : FClosure A f s x ) → FClosure A f s (f x)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
95
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
96 A∋fc : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A y
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
97 A∋fc {A} s f mf (init as) = as
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
98 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) )
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
99
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
100 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → * s ≤ * y
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
101 s≤fc {A} s {.s} f mf (init x) = case1 refl
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
102 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
103 ... | case1 x=fx = subst (λ k → * s ≤ * k ) (*≡*→≡ x=fx) ( s≤fc {A} s f mf fcy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
104 ... | case2 x<fx with s≤fc {A} s f mf fcy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
105 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym s≡x) refl x<fx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
106 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx )
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
107
557
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
108 fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
109 fcn s mf (init as) = zero
558
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
110 fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
111 ... | case1 eq = fcn s mf p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
112 ... | case2 y<fy = suc (fcn s mf p )
557
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
113
558
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
114 fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
115 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx ≡ fcn s mf cy → * x ≡ * y
559
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
116 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
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
117 fc00 : (i j : ℕ ) → i ≡ j → {x y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → i ≡ fcn s mf cx → j ≡ fcn s mf cy → * x ≡ * y
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
118 fc00 zero zero refl (init _) (init x₁) i=x i=y = refl
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
119 fc00 zero zero refl (init as) (fsuc y cy) i=x i=y with proj1 (mf y (A∋fc s f mf cy ) )
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
120 ... | case1 y=fy = subst (λ k → * s ≡ k ) y=fy ( fc00 zero zero refl (init as) cy i=x i=y )
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
121 fc00 zero zero refl (fsuc x cx) (init as) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) )
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
122 ... | case1 x=fx = subst (λ k → k ≡ * s ) x=fx ( fc00 zero zero refl cx (init as) i=x i=y )
559
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
123 fc00 zero zero refl (fsuc x cx) (fsuc y cy) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) )
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
124 ... | case1 x=fx | case1 y=fy = subst₂ (λ j k → j ≡ k ) x=fx y=fy ( fc00 zero zero refl cx cy i=x i=y )
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
125 fc00 (suc i) (suc j) i=j {.(f x)} {.(f y)} (fsuc x cx) (fsuc y cy) i=x j=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) )
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
126 ... | case1 x=fx | case1 y=fy = subst₂ (λ j k → j ≡ k ) x=fx y=fy ( fc00 (suc i) (suc j) i=j cx cy i=x j=y )
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
127 ... | case1 x=fx | case2 y<fy = subst (λ k → k ≡ * (f y)) x=fx (fc02 x cx i=x) where
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
128 fc02 : (x1 : Ordinal) → (cx1 : FClosure A f s x1 ) → suc i ≡ fcn s mf cx1 → * x1 ≡ * (f y)
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
129 fc02 .(f x1) (fsuc x1 cx1) i=x1 with proj1 (mf x1 (A∋fc s f mf cx1 ) )
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
130 ... | case1 eq = trans (sym eq) ( fc02 x1 cx1 i=x1 ) -- derefence while f x ≡ x
559
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
131 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc04) where
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
132 fc04 : * x1 ≡ * y
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
133 fc04 = fc00 i j (cong pred i=j) cx1 cy (cong pred i=x1) (cong pred j=y)
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
134 ... | case2 x<fx | case1 y=fy = subst (λ k → * (f x) ≡ k ) y=fy (fc03 y cy j=y) where
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
135 fc03 : (y1 : Ordinal) → (cy1 : FClosure A f s y1 ) → suc j ≡ fcn s mf cy1 → * (f x) ≡ * y1
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
136 fc03 .(f y1) (fsuc y1 cy1) j=y1 with proj1 (mf y1 (A∋fc s f mf cy1 ) )
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
137 ... | case1 eq = trans ( fc03 y1 cy1 j=y1 ) eq
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
138 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
139 fc05 : * x ≡ * y1
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
140 fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1)
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
141 ... | 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)))
557
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
142
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
143
557
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
144 fcn-< : {A : HOD} (s : Ordinal ) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
145 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx Data.Nat.< fcn s mf cy → * x < * y
558
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
146 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
147 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
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
148 fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
149 ... | case1 y=fy = subst (λ k → * x < k ) y=fy ( fc01 (suc i) {y1} cx cy i=y (s≤s x<i) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
150 ... | case2 y<fy with <-cmp (fcn s mf cx ) i
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
151 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> x<i c )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
152 ... | tri≈ ¬a b ¬c = subst (λ k → k < * (f y1) ) (fcn-inject s mf cy cx (sym (trans b (cong pred i=y) ))) y<fy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
153 ... | tri< a ¬b ¬c = IsStrictPartialOrder.trans PO fc02 y<fy where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
154 fc03 : suc i ≡ suc (fcn s mf cy) → i ≡ fcn s mf cy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
155 fc03 eq = cong pred eq
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
156 fc02 : * x < * y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 557
diff changeset
157 fc02 = fc01 i cx cy (fc03 i=y ) a
557
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
158
559
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
159 fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
160 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → Tri (* x < * y) (* x ≡ * y) (* y < * x )
559
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
161 fcn-cmp {A} s {x} {y} f mf cx cy with <-cmp ( fcn s mf cx ) (fcn s mf cy )
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
162 ... | tri< a ¬b ¬c = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
163 fc11 : * x < * y
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
164 fc11 = fcn-< {A} s {x} {y} {f} mf cx cy a
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
165 ... | tri≈ ¬a b ¬c = tri≈ (λ lt → <-irr (case1 (sym fc10)) lt) fc10 (λ lt → <-irr (case1 fc10) lt) where
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
166 fc10 : * x ≡ * y
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
167 fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
168 ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
169 fc12 : * y < * x
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
170 fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c
9ba98ecfbe62 fcn-cmp done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 558
diff changeset
171
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
172
562
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
173 fcn-imm : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
174 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → ¬ ( ( * x < * y ) ∧ ( * y < * (f x )) )
563
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
175 fcn-imm {A} s {x} {y} f mf cx cy ⟪ x<y , y<fx ⟫ = fc21 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
176 fc20 : fcn s mf cy Data.Nat.< suc (fcn s mf cx) → (fcn s mf cy ≡ fcn s mf cx) ∨ ( fcn s mf cy Data.Nat.< fcn s mf cx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
177 fc20 y<sx with <-cmp ( fcn s mf cy ) (fcn s mf cx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
178 ... | tri< a ¬b ¬c = case2 a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
179 ... | tri≈ ¬a b ¬c = case1 b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
180 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> y<sx (s≤s c))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
181 fc17 : {x y : Ordinal } → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → suc (fcn s mf cx) ≡ fcn s mf cy → * (f x ) ≡ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
182 fc17 {x} {y} cx cy sx=y = fc18 (fcn s mf cy) cx cy refl sx=y where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
183 fc18 : (i : ℕ ) → {y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → (i ≡ fcn s mf cy ) → suc (fcn s mf cx) ≡ i → * (f x) ≡ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
184 fc18 (suc i) {y} cx (fsuc y1 cy) i=y sx=i with proj1 (mf y1 (A∋fc s f mf cy ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
185 ... | case1 y=fy = subst (λ k → * (f x) ≡ k ) y=fy ( fc18 (suc i) {y1} cx cy i=y sx=i) -- dereference
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
186 ... | case2 y<fy = subst₂ (λ j k → * (f j) ≡ * (f k) ) &iso &iso (cong (λ k → * (f (& k) ) ) fc19) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
187 fc19 : * x ≡ * y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
188 fc19 = fcn-inject s mf cx cy (cong pred ( trans sx=i i=y ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
189 fc21 : ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
190 fc21 with <-cmp (suc ( fcn s mf cx )) (fcn s mf cy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
191 ... | tri< a ¬b ¬c = <-irr (case2 y<fx) (fc22 a) where -- suc ncx < ncy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
192 cxx : FClosure A f s (f x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
193 cxx = fsuc x cx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
194 fc16 : (x : Ordinal ) → (cx : FClosure A f s x) → (fcn s mf cx ≡ fcn s mf (fsuc x cx)) ∨ ( suc (fcn s mf cx ) ≡ fcn s mf (fsuc x cx))
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
195 fc16 x (init as) with proj1 (mf s as )
563
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
196 ... | case1 _ = case1 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
197 ... | case2 _ = case2 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
198 fc16 .(f x) (fsuc x cx ) with proj1 (mf (f x) (A∋fc s f mf (fsuc x cx)) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
199 ... | case1 _ = case1 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
200 ... | case2 _ = case2 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
201 fc22 : (suc ( fcn s mf cx )) Data.Nat.< (fcn s mf cy ) → * (f x) < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
202 fc22 a with fc16 x cx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
203 ... | case1 eq = fcn-< s mf cxx cy (subst (λ k → k Data.Nat.< fcn s mf cy ) eq (<-trans a<sa a))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
204 ... | case2 eq = fcn-< s mf cxx cy (subst (λ k → k Data.Nat.< fcn s mf cy ) eq a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
205 ... | tri≈ ¬a b ¬c = <-irr (case1 (fc17 cx cy b)) y<fx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
206 ... | tri> ¬a ¬b c with fc20 c -- ncy < suc ncx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
207 ... | case1 y=x = <-irr (case1 ( fcn-inject s mf cy cx y=x )) x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
208 ... | case2 y<x = <-irr (case2 x<y) (fcn-< s mf cy cx y<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
209
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
210 -- open import Relation.Binary.Properties.Poset as Poset
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
211
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
212 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
213 IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b) → Tri (a < b) (a ≡ b) (b < a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
214
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
215 ⊆-IsTotalOrderSet : { A B : HOD } → B ⊆ A → IsTotalOrderSet A → IsTotalOrderSet B
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
216 ⊆-IsTotalOrderSet {A} {B} B⊆A T ax ay = T (incl B⊆A ax) (incl B⊆A ay)
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
217
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
218 _⊆'_ : ( A B : HOD ) → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
219 _⊆'_ A B = {x : Ordinal } → odef A x → odef B x
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
220
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
221 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
222 -- inductive maxmum tree from x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
223 -- tree structure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
224 --
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
225
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
226 record HasPrev (A B : HOD) {x : Ordinal } (xa : odef A x) ( f : Ordinal → Ordinal ) : Set n where
533
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
227 field
534
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 533
diff changeset
228 y : Ordinal
541
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 540
diff changeset
229 ay : odef B y
534
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 533
diff changeset
230 x=fy : x ≡ f y
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
231
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
232 record IsSup (A B : HOD) {x : Ordinal } (xa : odef A x) : Set n where
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
233 field
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
234 x<sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x )
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
235
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
236 record ZChain ( A : HOD ) (x : Ordinal) ( f : Ordinal → Ordinal ) (supf : Ordinal → HOD) ( z : Ordinal ) : Set (Level.suc n) where
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
237 chain : HOD
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
238 chain = supf z
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
239 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
240 chain⊆A : chain ⊆' A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
241 chain∋x : odef chain x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
242 initial : {y : Ordinal } → odef chain y → * x ≤ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
243 f-next : {a : Ordinal } → odef chain a → odef chain (f a)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
244 f-immediate : { x y : Ordinal } → odef chain x → odef chain y → ¬ ( ( * x < * y ) ∧ ( * y < * (f x )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
245 is-max : {a b : Ordinal } → (ca : odef chain a ) → b o< osuc z → (ab : odef A b)
574
9084a26445a7 ZC data won't work
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 573
diff changeset
246 → HasPrev A chain ab f ∨ IsSup A chain ab
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
247 → * a < * b → odef chain b
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
248
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
249 record ZChain1 ( A : HOD ) (x : Ordinal) ( f : Ordinal → Ordinal ) ( z : Ordinal ) : Set (Level.suc n) where
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
250 field
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
251 supf : Ordinal → HOD
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
252 zc : ZChain A x f supf z
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
253 chain-mono : {x y : Ordinal} → x o≤ y → y o≤ z → supf x ⊆' supf y
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
254 f-total : {x y : Ordinal} → x o≤ z → IsTotalOrderSet (supf x)
595
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 594
diff changeset
255
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
256 record Maximal ( A : HOD ) : Set (Level.suc n) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
257 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
258 maximal : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
259 A∋maximal : A ∋ maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
260 ¬maximal<x : {x : HOD} → A ∋ x → ¬ maximal < x -- A is Partial, use negative
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
261
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
262 record SUP ( A B : HOD ) : Set (Level.suc n) where
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
263 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
264 sup : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
265 A∋maximal : A ∋ sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
266 x<sup : {x : HOD} → B ∋ x → (x ≡ sup ) ∨ (x < sup ) -- B is Total, use positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
267
533
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
268 SupCond : ( A B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → Set (Level.suc n)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
269 SupCond A B _ _ = SUP A B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
270
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
271 Zorn-lemma : { A : HOD }
464
5acf6483a9e3 Zorn lemma start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 447
diff changeset
272 → o∅ o< & A
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
273 → ( ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
274 → Maximal A
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
275 Zorn-lemma {A} 0<A supP = zorn00 where
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
276 supO : (C : HOD ) → C ⊆' A → IsTotalOrderSet C → Ordinal
566
a64dad8d2e93 fix sp1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 565
diff changeset
277 supO C C⊆A TC = & ( SUP.sup ( supP C C⊆A TC ))
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
278 <-irr0 : {a b : HOD} → A ∋ a → A ∋ b → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
279 <-irr0 {a} {b} A∋a A∋b = <-irr
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
280 z07 : {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
281 z07 {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
282 s : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
283 s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A ))
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
284 as : A ∋ * ( & s )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
285 as = subst (λ k → odef A (& k) ) (sym *iso) ( ODC.x∋minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) )
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
286 as0 : odef A (& s )
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
287 as0 = subst (λ k → odef A k ) &iso as
547
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 546
diff changeset
288 s<A : & s o< & A
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
289 s<A = c<→o< (subst (λ k → odef A (& k) ) *iso as )
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
290 HasMaximal : HOD
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
291 HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) → odef A m → ¬ (* x < * m)) } ; odmax = & A ; <odmax = z07 }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
292 no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → odef A x ∧ ((m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m) )) → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
293 no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
294 Gtx : { x : HOD} → A ∋ x → HOD
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
295 Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
296 z08 : ¬ Maximal A → HasMaximal =h= od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
297 z08 nmx = record { eq→ = λ {x} lt → ⊥-elim ( nmx record {maximal = * x ; A∋maximal = subst (λ k → odef A k) (sym &iso) (proj1 lt)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
298 ; ¬maximal<x = λ {y} ay → subst (λ k → ¬ (* x < k)) *iso (proj2 lt (& y) ay) } ) ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
299 x-is-maximal : ¬ Maximal A → {x : Ordinal} → (ax : odef A x) → & (Gtx (subst (λ k → odef A k ) (sym &iso) ax)) ≡ o∅ → (m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
300 x-is-maximal nmx {x} ax nogt m am = ⟪ subst (λ k → odef A k) &iso (subst (λ k → odef A k ) (sym &iso) ax) , ¬x<m ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
301 ¬x<m : ¬ (* x < * m)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
302 ¬x<m x<m = ∅< {Gtx (subst (λ k → odef A k ) (sym &iso) ax)} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫ (≡o∅→=od∅ nogt)
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
303
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
304 -- Uncountable ascending chain by axiom of choice
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
305 cf : ¬ Maximal A → Ordinal → Ordinal
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
306 cf nmx x with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
307 ... | no _ = o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
308 ... | yes ax with is-o∅ (& ( Gtx ax ))
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
309 ... | yes nogt = -- no larger element, so it is maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
310 ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
311 ... | no not = & (ODC.minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq)))
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
312 is-cf : (nmx : ¬ Maximal A ) → {x : Ordinal} → odef A x → odef A (cf nmx x) ∧ ( * x < * (cf nmx x) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
313 is-cf nmx {x} ax with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
314 ... | no not = ⊥-elim ( not (subst (λ k → odef A k ) (sym &iso) ax ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
315 ... | yes ax with is-o∅ (& ( Gtx ax ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
316 ... | yes nogt = ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
317 ... | no not = ODC.x∋minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq))
606
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
318
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
319 ---
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
320 --- infintie ascention sequence of f
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
321 ---
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
322 cf-is-<-monotonic : (nmx : ¬ Maximal A ) → (x : Ordinal) → odef A x → ( * x < * (cf nmx x) ) ∧ odef A (cf nmx x )
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
323 cf-is-<-monotonic nmx x ax = ⟪ proj2 (is-cf nmx ax ) , proj1 (is-cf nmx ax ) ⟫
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
324 cf-is-≤-monotonic : (nmx : ¬ Maximal A ) → ≤-monotonic-f A ( cf nmx )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
325 cf-is-≤-monotonic nmx x ax = ⟪ case2 (proj1 ( cf-is-<-monotonic nmx x ax )) , proj2 ( cf-is-<-monotonic nmx x ax ) ⟫
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
326
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
327 zsup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f) → (zc1 : ZChain1 A (& s) f (& A) ) → SUP A (ZChain.chain (ZChain1.zc zc1))
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
328 zsup f mf zc1 = supP (ZChain.chain zc) (ZChain.chain⊆A zc) ( ZChain1.f-total zc1 {& A} {& A} o≤-refl ) where
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
329 zc = ZChain1.zc zc1
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
330 A∋zsup : (nmx : ¬ Maximal A ) (zc1 : ZChain1 A (& s) (cf nmx) (& A) )
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
331 → A ∋ * ( & ( SUP.sup (zsup (cf nmx) (cf-is-≤-monotonic nmx) zc1 )))
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
332 A∋zsup nmx zc1 = subst (λ k → odef A (& k )) (sym *iso) ( SUP.A∋maximal (zsup (cf nmx) (cf-is-≤-monotonic nmx) zc1 ) )
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
333 sp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (zc1 : ZChain1 A (& s) f (& A) ) → SUP A (ZChain.chain (ZChain1.zc zc1))
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
334 sp0 f mf zc1 = supP (ZChain.chain zc) (ZChain.chain⊆A zc) (ZChain1.f-total zc1 {& A} {& A} o≤-refl ) where
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
335 zc = ZChain1.zc zc1
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
336 zc< : {x y z : Ordinal} → {P : Set n} → (x o< y → P) → x o< z → z o< y → P
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
337 zc< {x} {y} {z} {P} prev x<z z<y = prev (ordtrans x<z z<y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
338
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
339 ---
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
340 --- the maximum chain has fix point of any ≤-monotonic function
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
341 ---
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
342 fixpoint : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (zc1 : ZChain1 A (& s) f (& A) )
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
343 → f (& (SUP.sup (sp0 f mf zc1 ))) ≡ & (SUP.sup (sp0 f mf zc1 ))
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
344 fixpoint f mf zc1 = z14 where
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
345 zc = ZChain1.zc zc1
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
346 chain = ZChain.chain zc
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
347 sp1 = sp0 f mf zc1
565
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 564
diff changeset
348 z10 : {a b : Ordinal } → (ca : odef chain a ) → b o< osuc (& A) → (ab : odef A b )
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
349 → HasPrev A chain ab f ∨ IsSup A chain {b} ab -- (supO chain (ZChain.chain⊆A zc) (ZChain.f-total zc) ≡ b )
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
350 → * a < * b → odef chain b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
351 z10 = ZChain.is-max zc
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
352 z11 : & (SUP.sup sp1) o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
353 z11 = c<→o< ( SUP.A∋maximal sp1)
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
354 z12 : odef chain (& (SUP.sup sp1))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
355 z12 with o≡? (& s) (& (SUP.sup sp1))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
356 ... | yes eq = subst (λ k → odef chain k) eq ( ZChain.chain∋x zc )
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
357 ... | no ne = z10 {& s} {& (SUP.sup sp1)} ( ZChain.chain∋x zc ) (ordtrans z11 <-osuc ) (SUP.A∋maximal sp1)
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
358 (case2 z19 ) z13 where
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
359 z13 : * (& s) < * (& (SUP.sup sp1))
566
a64dad8d2e93 fix sp1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 565
diff changeset
360 z13 with SUP.x<sup sp1 ( ZChain.chain∋x zc )
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
361 ... | case1 eq = ⊥-elim ( ne (cong (&) eq) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
362 ... | case2 lt = subst₂ (λ j k → j < k ) (sym *iso) (sym *iso) lt
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
363 z19 : IsSup A chain {& (SUP.sup sp1)} (SUP.A∋maximal sp1)
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
364 z19 = record { x<sup = z20 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
365 z20 : {y : Ordinal} → odef chain y → (y ≡ & (SUP.sup sp1)) ∨ (y << & (SUP.sup sp1))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
366 z20 {y} zy with SUP.x<sup sp1 (subst (λ k → odef chain k ) (sym &iso) zy)
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
367 ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso ( cong (&) y=p ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
368 ... | case2 y<p = case2 (subst (λ k → * y < k ) (sym *iso) y<p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
369 -- λ {y} zy → subst (λ k → (y ≡ & k ) ∨ (y << & k)) ? (SUP.x<sup sp1 ? ) }
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
370 z14 : f (& (SUP.sup (sp0 f mf zc1))) ≡ & (SUP.sup (sp0 f mf zc1))
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
371 z14 with ZChain1.f-total zc1 {& A} {& A} o≤-refl (subst (λ k → odef chain k) (sym &iso) (ZChain.f-next zc z12 )) z12
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
372 ... | tri< a ¬b ¬c = ⊥-elim z16 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
373 z16 : ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
374 z16 with proj1 (mf (& ( SUP.sup sp1)) ( SUP.A∋maximal sp1 ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
375 ... | case1 eq = ⊥-elim (¬b (subst₂ (λ j k → j ≡ k ) refl *iso (sym eq) ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
376 ... | case2 lt = ⊥-elim (¬c (subst₂ (λ j k → k < j ) refl *iso lt ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
377 ... | tri≈ ¬a b ¬c = subst ( λ k → k ≡ & (SUP.sup sp1) ) &iso ( cong (&) b )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
378 ... | tri> ¬a ¬b c = ⊥-elim z17 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
379 z15 : (* (f ( & ( SUP.sup sp1 ))) ≡ SUP.sup sp1) ∨ (* (f ( & ( SUP.sup sp1 ))) < SUP.sup sp1)
566
a64dad8d2e93 fix sp1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 565
diff changeset
380 z15 = SUP.x<sup sp1 (subst (λ k → odef chain k ) (sym &iso) (ZChain.f-next zc z12 ))
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
381 z17 : ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
382 z17 with z15
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
383 ... | case1 eq = ¬b eq
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
384 ... | case2 lt = ¬a lt
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
385
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
386 -- ZChain contradicts ¬ Maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
387 --
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
388 -- ZChain forces fix point on any ≤-monotonic function (fixpoint)
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
389 -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
390 --
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
391 z04 : (nmx : ¬ Maximal A ) → (zc : ZChain1 A (& s) (cf nmx) (& A)) → ⊥
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
392 z04 nmx zc1 = <-irr0 {* (cf nmx c)} {* c} (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (SUP.A∋maximal sp1))))
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
393 (subst (λ k → odef A (& k)) (sym *iso) (SUP.A∋maximal sp1) )
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
394 (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) zc1 ))) -- x ≡ f x ̄
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
395 (proj1 (cf-is-<-monotonic nmx c (SUP.A∋maximal sp1))) where -- x < f x
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
396 zc = ZChain1.zc zc1
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
397 sp1 = sp0 (cf nmx) (cf-is-≤-monotonic nmx) zc1
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
398 c = & (SUP.sup sp1)
548
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 547
diff changeset
399
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
400 --
547
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 546
diff changeset
401 -- create all ZChains under o< x
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
402 --
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
403
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
404 ind : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → (x : Ordinal) → ((y : Ordinal) → y o< x → { y₁ : Ordinal} (ay : odef A y₁)
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
405 → ZChain1 A y₁ f y) → {y : Ordinal} (ay : odef A y) → ZChain1 A y f x
547
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 546
diff changeset
406 ind f mf x prev {y} ay with Oprev-p x
548
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 547
diff changeset
407 ... | yes op = zc4 where
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
408 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
409 -- we have previous ordinal to use induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
410 --
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
411 open ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
412
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
413 px = Oprev.oprev op
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
414 supf : Ordinal → HOD
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
415 supf = ZChain1.supf (prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) ay)
610
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
416 zc1 : ZChain1 A y f (Oprev.oprev op)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
417 zc1 = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) ay
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
418 zc0 : ZChain A y f (ZChain1.supf (prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) ay)) (Oprev.oprev op)
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
419 zc0 = ZChain1.zc (prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) ay)
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
420 zc0-b<x : (b : Ordinal ) → b o< x → b o< osuc px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
421 zc0-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
422 px<x : px o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
423 px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
424
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
425 -- if previous chain satisfies maximality, we caan reuse it
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
426 --
610
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
427 no-extenion : ( {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
428 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
429 * a < * b → odef (ZChain.chain zc0) b ) → ZChain1 A y f x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
430 no-extenion is-max = record { supf = supf0 ; zc = record { chain⊆A = subst (λ k → k ⊆' A ) seq (ZChain.chain⊆A zc0)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
431 ; initial = subst (λ k → {y₁ : Ordinal} → odef k y₁ → * y ≤ * y₁ ) seq (ZChain.initial zc0)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
432 ; f-next = subst (λ k → {a : Ordinal} → odef k a → odef k (f a) ) seq (ZChain.f-next zc0)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
433 ; f-immediate = subst (λ k → {x₁ : Ordinal} {y₁ : Ordinal} → odef k x₁ → odef k y₁ →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
434 ¬ (* x₁ < * y₁) ∧ (* y₁ < * (f x₁)) ) seq (ZChain.f-immediate zc0) ; chain∋x = subst (λ k → odef k y ) seq (ZChain.chain∋x zc0)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
435 ; is-max = subst (λ k → {a b : Ordinal} → odef k a → b o< osuc x → (ab : odef A b) →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
436 HasPrev A k ab f ∨ IsSup A k ab → * a < * b → odef k b ) seq is-max }
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
437 ; chain-mono = mono ; f-total = {!!} } where
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
438 supf0 : Ordinal → HOD
610
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
439 supf0 z with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
440 ... | tri< a ¬b ¬c = supf z
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
441 ... | tri≈ ¬a b ¬c = ZChain.chain zc0
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
442 ... | tri> ¬a ¬b c = ZChain.chain zc0
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
443 seq : ZChain.chain zc0 ≡ supf0 x
610
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
444 seq with trio< x x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
445 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b refl )
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
446 ... | tri≈ ¬a b ¬c = refl
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
447 ... | tri> ¬a ¬b c = refl
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
448 seq<x : {b : Ordinal } → b o< x → supf b ≡ supf0 b
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
449 seq<x {b} b<x with trio< b x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
450 ... | tri< a ¬b ¬c = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
451 ... | tri≈ ¬a b₁ ¬c = ⊥-elim (¬a b<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
452 ... | tri> ¬a ¬b c = ⊥-elim (¬a b<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
453 mono : {a b : Ordinal} → a o≤ b → b o< osuc x →
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
454 supf0 a ⊆' supf0 b
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
455 mono {a} {b} a≤b b<ox with osuc-≡< a≤b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
456 ... | case1 refl = λ x → x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
457 ... | case2 a<b with osuc-≡< b<ox
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
458 ... | case1 b=x = subst₂ (λ j k → j ⊆' k ) (seq<x a<x) nc00 ( ZChain1.chain-mono zc1 a≤px <-osuc ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
459 a<x : a o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
460 a<x with osuc-≡< b<ox
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
461 ... | case1 b=x = subst (λ k → a o< k ) b=x a<b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
462 ... | case2 b<x = ordtrans a<b b<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
463 a≤px : a o≤ px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
464 a≤px = subst (λ k → a o< k ) (sym (Oprev.oprev=x op)) a<x
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
465 nc00 : supf px ≡ supf0 b
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
466 nc00 with trio< b x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
467 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b b=x )
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
468 ... | tri≈ ¬a b ¬c = refl
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
469 ... | tri> ¬a ¬b c = ⊥-elim ( ¬b b=x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
470 ... | case2 b<x = subst₂ (λ j k → j ⊆' k ) (seq<x a<x ) (seq<x b<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
471 ( ZChain1.chain-mono zc1 a≤b (subst (λ k → b o< k) (sym (Oprev.oprev=x op)) b<x ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
472 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
473 a<x : a o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
474 a<x with osuc-≡< b<ox
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
475 ... | case1 b=x = subst (λ k → a o< k ) b=x a<b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
476 ... | case2 b<x = ordtrans a<b b<x
610
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 609
diff changeset
477
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
478 zc4 : ZChain1 A y f x
565
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 564
diff changeset
479 zc4 with ODC.∋-p O A (* x)
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
480 ... | no noax = no-extenion zc11 where -- ¬ A ∋ p, just skip
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
481 zc11 : {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) →
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
482 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab →
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
483 * a < * b → odef (ZChain.chain zc0) b
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
484 zc11 {a} {b} za b<ox ab P a<b with osuc-≡< b<ox
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
485 ... | case1 eq = ⊥-elim ( noax (subst (λ k → odef A k) (trans eq (sym &iso)) ab ) )
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
486 ... | case2 lt = ZChain.is-max zc0 za (zc0-b<x b lt) ab P a<b
582
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 580
diff changeset
487 ... | yes ax with ODC.p∨¬p O ( HasPrev A (ZChain.chain zc0) ax f ) -- we have to check adding x preserve is-max ZChain A y f mf supO x
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
488 ... | case1 pr = no-extenion zc17 where -- we have previous A ∋ z < x , f z ≡ x, so chain ∋ f z ≡ x because of f-next
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
489 chain0 = ZChain.chain zc0
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
490 zc17 : {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) →
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
491 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab →
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
492 * a < * b → odef (ZChain.chain zc0) b
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
493 zc17 {a} {b} za b<ox ab P a<b with osuc-≡< b<ox
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
494 ... | case2 lt = ZChain.is-max zc0 za (zc0-b<x b lt) ab P a<b
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
495 ... | case1 b=x = subst (λ k → odef chain0 k ) (trans (sym (HasPrev.x=fy pr )) (trans &iso (sym b=x)) ) ( ZChain.f-next zc0 (HasPrev.ay pr))
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
496 ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A (ZChain.chain zc0) ax )
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
497 ... | case1 is-sup = -- x is a sup of zc0
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
498 record { zc = record { chain⊆A = {!!} ; f-next = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
499 ; initial = {!!} ; f-immediate = {!!} ; chain∋x = {!!} ; is-max = {!!} } ; supf = supf0 ; chain-mono = mono ; f-total = {!!} } where
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
500 sup0 : SUP A (ZChain.chain zc0)
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
501 sup0 = record { sup = * x ; A∋maximal = ax ; x<sup = x21 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
502 x21 : {y : HOD} → ZChain.chain zc0 ∋ y → (y ≡ * x) ∨ (y < * x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
503 x21 {y} zy with IsSup.x<sup is-sup zy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
504 ... | case1 y=x = case1 ( subst₂ (λ j k → j ≡ * k ) *iso &iso ( cong (*) y=x) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
505 ... | case2 y<x = case2 (subst₂ (λ j k → j < * k ) *iso &iso y<x )
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
506 sp : HOD
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
507 sp = SUP.sup sup0
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
508 x=sup : x ≡ & sp
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
509 x=sup = sym &iso
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
510 chain0 = ZChain.chain zc0
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
511 sc<A : {y : Ordinal} → odef chain0 y ∨ FClosure A f (& sp) y → y o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
512 sc<A {y} (case1 zx) = subst (λ k → k o< (& A)) &iso ( c<→o< (ZChain.chain⊆A zc0 (subst (λ k → odef chain0 k) (sym &iso) zx )))
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
513 sc<A {y} (case2 fx) = subst (λ k → k o< (& A)) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso) (A∋fc (& sp) f mf fx )) )
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
514 schain : HOD
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
515 schain = record { od = record { def = λ x → odef chain0 x ∨ (FClosure A f (& sp) x) } ; odmax = & A ; <odmax = λ {y} sy → sc<A {y} sy }
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
516 A∋schain : {x : HOD } → schain ∋ x → A ∋ x
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
517 A∋schain (case1 zx ) = ZChain.chain⊆A zc0 zx
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
518 A∋schain {y} (case2 fx ) = A∋fc (& sp) f mf fx
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
519 s⊆A : schain ⊆' A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
520 s⊆A {x} (case1 zx) = ZChain.chain⊆A zc0 zx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
521 s⊆A {x} (case2 fx) = A∋fc (& sp) f mf fx
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
522 cmp : {a b : HOD} (za : odef chain0 (& a)) (fb : FClosure A f (& sp) (& b)) → Tri (a < b) (a ≡ b) (b < a )
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
523 cmp {a} {b} za fb with SUP.x<sup sup0 za | s≤fc (& sp) f mf fb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
524 ... | case1 sp=a | case1 sp=b = tri≈ (λ lt → <-irr (case1 (sym eq)) lt ) eq (λ lt → <-irr (case1 eq) lt ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
525 eq : a ≡ b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
526 eq = trans sp=a (subst₂ (λ j k → j ≡ k ) *iso *iso sp=b )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
527 ... | case1 sp=a | case2 sp<b = tri< a<b (λ eq → <-irr (case1 (sym eq)) a<b ) (λ lt → <-irr (case2 a<b) lt ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
528 a<b : a < b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
529 a<b = subst (λ k → k < b ) (sym sp=a) (subst₂ (λ j k → j < k ) *iso *iso sp<b )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
530 ... | case2 a<sp | case1 sp=b = tri< a<b (λ eq → <-irr (case1 (sym eq)) a<b ) (λ lt → <-irr (case2 a<b) lt ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
531 a<b : a < b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
532 a<b = subst (λ k → a < k ) (trans sp=b *iso ) (subst (λ k → a < k ) (sym *iso) a<sp )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
533 ... | case2 a<sp | case2 sp<b = tri< a<b (λ eq → <-irr (case1 (sym eq)) a<b ) (λ lt → <-irr (case2 a<b) lt ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
534 a<b : a < b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
535 a<b = ptrans (subst (λ k → a < k ) (sym *iso) a<sp ) ( subst₂ (λ j k → j < k ) refl *iso sp<b )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
536 scmp : {a b : HOD} → odef schain (& a) → odef schain (& b) → Tri (a < b) (a ≡ b) (b < a )
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
537 scmp {a} {b} (case1 za) (case1 zb) = ZChain1.f-total zc1 {px} {px} o≤-refl za zb
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
538 scmp {a} {b} (case1 za) (case2 fb) = cmp za fb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
539 scmp (case2 fa) (case1 zb) with cmp zb fa
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
540 ... | tri< a ¬b ¬c = tri> ¬c (λ eq → ¬b (sym eq)) a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
541 ... | tri≈ ¬a b ¬c = tri≈ ¬c (sym b) ¬a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
542 ... | tri> ¬a ¬b c = tri< c (λ eq → ¬b (sym eq)) ¬a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
543 scmp (case2 fa) (case2 fb) = subst₂ (λ a b → Tri (a < b) (a ≡ b) (b < a ) ) *iso *iso (fcn-cmp (& sp) f mf fa fb)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
544 scnext : {a : Ordinal} → odef schain a → odef schain (f a)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
545 scnext {x} (case1 zx) = case1 (ZChain.f-next zc0 zx)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
546 scnext {x} (case2 sx) = case2 ( fsuc x sx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
547 scinit : {x : Ordinal} → odef schain x → * y ≤ * x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
548 scinit {x} (case1 zx) = ZChain.initial zc0 zx
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
549 scinit {x} (case2 sx) with (s≤fc (& sp) f mf sx ) | SUP.x<sup sup0 (subst (λ k → odef chain0 k ) (sym &iso) ( ZChain.chain∋x zc0 ) )
562
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
550 ... | case1 sp=x | case1 y=sp = case1 (trans y=sp (subst (λ k → k ≡ * x ) *iso sp=x ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
551 ... | case1 sp=x | case2 y<sp = case2 (subst (λ k → * y < k ) (trans (sym *iso) sp=x) y<sp )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
552 ... | case2 sp<x | case1 y=sp = case2 (subst (λ k → k < * x ) (trans *iso (sym y=sp )) sp<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
553 ... | case2 sp<x | case2 y<sp = case2 (ptrans y<sp (subst (λ k → k < * x ) *iso sp<x) )
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
554 A∋za : {a : Ordinal } → odef chain0 a → odef A a
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
555 A∋za za = ZChain.chain⊆A zc0 za
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
556 za<sup : {a : Ordinal } → odef chain0 a → ( * a ≡ sp ) ∨ ( * a < sp )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
557 za<sup za = SUP.x<sup sup0 (subst (λ k → odef chain0 k ) (sym &iso) za )
562
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
558 simm : {a b : Ordinal} → odef schain a → odef schain b → ¬ (* a < * b) ∧ (* b < * (f a))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
559 simm {a} {b} (case1 za) (case1 zb) = ZChain.f-immediate zc0 za zb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
560 simm {a} {b} (case1 za) (case2 sb) p with proj1 (mf a (A∋za za) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
561 ... | case1 eq = <-irr (case2 (subst (λ k → * b < k ) (sym eq) (proj2 p))) (proj1 p)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
562 ... | case2 a<fa with za<sup ( ZChain.f-next zc0 za ) | s≤fc (& sp) f mf sb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
563 ... | case1 fa=sp | case1 sp=b = <-irr (case1 (trans fa=sp (trans (sym *iso) sp=b )) ) ( proj2 p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
564 ... | case2 fa<sp | case1 sp=b = <-irr (case2 fa<sp) (subst (λ k → k < * (f a) ) (trans (sym sp=b) *iso) (proj2 p ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
565 ... | case1 fa=sp | case2 sp<b = <-irr (case2 (proj2 p )) (subst (λ k → k < * b) (sym fa=sp) (subst (λ k → k < * b ) *iso sp<b ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
566 ... | case2 fa<sp | case2 sp<b = <-irr (case2 (proj2 p )) (ptrans fa<sp (subst (λ k → k < * b ) *iso sp<b ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
567 simm {a} {b} (case2 sa) (case1 zb) p with proj1 (mf a ( subst (λ k → odef A k) &iso ( A∋schain (case2 (subst (λ k → FClosure A f (& sp) k ) (sym &iso) sa) )) ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
568 ... | case1 eq = <-irr (case2 (subst (λ k → * b < k ) (sym eq) (proj2 p))) (proj1 p)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
569 ... | case2 b<fb with s≤fc (& sp) f mf sa | za<sup zb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
570 ... | case1 sp=a | case1 b=sp = <-irr (case1 (trans b=sp (trans (sym *iso) sp=a )) ) (proj1 p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
571 ... | case1 sp=a | case2 b<sp = <-irr (case2 (subst (λ k → * b < k ) (trans (sym *iso) sp=a) b<sp ) ) (proj1 p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
572 ... | case2 sp<a | case1 b=sp = <-irr (case2 (subst ( λ k → k < * a ) (trans *iso (sym b=sp)) sp<a )) (proj1 p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
573 ... | case2 sp<a | case2 b<sp = <-irr (case2 (ptrans b<sp (subst (λ k → k < * a) *iso sp<a ))) (proj1 p )
564
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 563
diff changeset
574 simm {a} {b} (case2 sa) (case2 sb) p = fcn-imm {A} (& sp) {a} {b} f mf sa sb p
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
575 s-ismax : {a b : Ordinal} → odef schain a → b o< osuc x → (ab : odef A b)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
576 → HasPrev A schain ab f ∨ IsSup A schain ab
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
577 → * a < * b → odef schain b
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
578 s-ismax {a} {b} sa b<ox ab p a<b with osuc-≡< b<ox -- b is x?
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
579 ... | case1 b=x = case2 (subst (λ k → FClosure A f (& sp) k ) (sym (trans b=x x=sup )) (init (SUP.A∋maximal sup0) ))
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
580 s-ismax {a} {b} (case1 za) b<ox ab (case1 p) a<b | case2 b<x = z21 p where -- has previous
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
581 z21 : HasPrev A schain ab f → odef schain b
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
582 z21 record { y = y ; ay = (case1 zy) ; x=fy = x=fy } =
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
583 case1 (ZChain.is-max zc0 za (zc0-b<x b b<x) ab (case1 record { y = y ; ay = zy ; x=fy = x=fy }) a<b )
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
584 z21 record { y = y ; ay = (case2 sy) ; x=fy = x=fy } = subst (λ k → odef schain k) (sym x=fy) (case2 (fsuc y sy) )
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
585 s-ismax {a} {b} (case1 za) b<ox ab (case2 p) a<b | case2 b<x = case1 (ZChain.is-max zc0 za (zc0-b<x b b<x) ab (case2 z22) a<b ) where -- previous sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
586 z22 : IsSup A (ZChain.chain zc0) ab
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
587 z22 = record { x<sup = λ {y} zy → IsSup.x<sup p (case1 zy ) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
588 s-ismax {a} {b} (case2 sa) b<ox ab (case1 p) a<b | case2 b<x with HasPrev.ay p
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
589 ... | case1 zy = case1 (subst (λ k → odef chain0 k ) (sym (HasPrev.x=fy p)) (ZChain.f-next zc0 zy )) -- in previous closure of f
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
590 ... | case2 sy = case2 (subst (λ k → FClosure A f (& (* x)) k ) (sym (HasPrev.x=fy p)) (fsuc (HasPrev.y p) sy )) -- in current closure of f
572
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
591 s-ismax {a} {b} (case2 sa) b<ox ab (case2 p) a<b | case2 b<x = case1 z23 where -- sup o< x is already in zc0
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
592 z24 : IsSup A schain ab → IsSup A (ZChain.chain zc0) ab
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
593 z24 p = record { x<sup = λ {y} zy → IsSup.x<sup p (case1 zy ) }
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
594 z23 : odef chain0 b
572
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
595 z23 with IsSup.x<sup (z24 p) ( ZChain.chain∋x zc0 )
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
596 ... | case1 y=b = subst (λ k → odef chain0 k ) y=b ( ZChain.chain∋x zc0 )
572
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
597 ... | case2 y<b = ZChain.is-max zc0 (ZChain.chain∋x zc0 ) (zc0-b<x b b<x) ab (case2 (z24 p)) y<b
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
598 supf0 : Ordinal → HOD
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
599 supf0 z with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
600 ... | tri< a ¬b ¬c = supf z
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
601 ... | tri≈ ¬a b ¬c = schain
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
602 ... | tri> ¬a ¬b c = schain
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
603 seq : schain ≡ supf0 x
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
604 seq with trio< x x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
605 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b refl )
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
606 ... | tri≈ ¬a b ¬c = refl
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
607 ... | tri> ¬a ¬b c = refl
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
608 seq<x : {b : Ordinal } → b o< x → supf b ≡ supf0 b
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
609 seq<x {b} b<x with trio< b x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
610 ... | tri< a ¬b ¬c = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
611 ... | tri≈ ¬a b₁ ¬c = ⊥-elim (¬a b<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
612 ... | tri> ¬a ¬b c = ⊥-elim (¬a b<x )
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
613 mono : {a b : Ordinal} → a o≤ b → b o< osuc x → supf0 a ⊆' supf0 b
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
614 mono {a} {b} a≤b b<ox = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
615
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
616 ... | case2 ¬x=sup = no-extenion z18 where -- x is not f y' nor sup of former ZChain from y -- no extention
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
617 z18 : {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) →
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
618 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab →
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
619 * a < * b → odef (ZChain.chain zc0) b
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
620 z18 {a} {b} za b<x ab p a<b with osuc-≡< b<x
569
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 568
diff changeset
621 ... | case2 lt = ZChain.is-max zc0 za (zc0-b<x b lt) ab p a<b
565
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 564
diff changeset
622 ... | case1 b=x with p
567
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 566
diff changeset
623 ... | case1 pr = ⊥-elim ( ¬fy<x record {y = HasPrev.y pr ; ay = HasPrev.ay pr ; x=fy = trans (trans &iso (sym b=x) ) (HasPrev.x=fy pr ) } )
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
624 ... | case2 b=sup = ⊥-elim ( ¬x=sup record {
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
625 x<sup = λ {y} zy → subst (λ k → (y ≡ k) ∨ (y << k)) (trans b=x (sym &iso)) (IsSup.x<sup b=sup zy) } )
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
626 ... | no ¬ox = record { supf = supf0 ; chain-mono = {!!} ; f-total = u-total
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
627 ; zc = record { chain⊆A = {!!} ; f-next = {!!}
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
628 ; initial = {!!} ; f-immediate = {!!} ; chain∋x = {!!} ; is-max = {!!} } } where --- limit ordinal case
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
629 record UZFChain (z : Ordinal) : Set n where -- Union of ZFChain from y which has maximality o< x
553
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 552
diff changeset
630 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 552
diff changeset
631 u : Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 552
diff changeset
632 u<x : u o< x
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
633 chain∋z : odef (ZChain.chain (ZChain1.zc (prev u u<x {y} ay ))) z
572
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
634 Uz⊆A : {z : Ordinal} → UZFChain z → odef A z
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
635 Uz⊆A {z} u = ZChain.chain⊆A (ZChain1.zc ( prev (UZFChain.u u) (UZFChain.u<x u) {y} ay )) (UZFChain.chain∋z u)
609
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 608
diff changeset
636 uzc1 : {z : Ordinal} → (u : UZFChain z) → ZChain1 A y f (UZFChain.u u)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 608
diff changeset
637 uzc1 {z} u = prev (UZFChain.u u) (UZFChain.u<x u) {y} ay
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 608
diff changeset
638 uzc : {z : Ordinal} → (u : UZFChain z) → ZChain A y f (ZChain1.supf (uzc1 u)) (UZFChain.u u)
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
639 uzc {z} u = ZChain1.zc (prev (UZFChain.u u) (UZFChain.u<x u) {y} ay)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
640 Uz : HOD
572
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
641 Uz = record { od = record { def = λ y → UZFChain y } ; odmax = & A
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
642 ; <odmax = λ lt → subst (λ k → k o< & A ) &iso (c<→o< (subst (λ k → odef A k ) (sym &iso) (Uz⊆A lt))) }
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
643 u-next : {z : Ordinal} → odef Uz z → odef Uz (f z)
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
644 u-next {z} u = record { u = UZFChain.u u ; u<x = UZFChain.u<x u ; chain∋z = ZChain.f-next ( uzc u ) (UZFChain.chain∋z u) }
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
645 u-initial : {z : Ordinal} → odef Uz z → * y ≤ * z
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
646 u-initial {z} u = ZChain.initial ( uzc u ) (UZFChain.chain∋z u)
427e36467a18 ZChain is monotonic on x, should be in record ZFChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 571
diff changeset
647 u-chain∋x : odef Uz y
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
648 u-chain∋x = record { u = y ; u<x = {!!} ; chain∋z = ZChain.chain∋x (ZChain1.zc (prev y {!!} ay )) }
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
649 supf0 : Ordinal → HOD
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
650 supf0 z with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
651 ... | tri< a ¬b ¬c = ZChain1.supf (prev z a {y} ay) z
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
652 ... | tri≈ ¬a b ¬c = Uz
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
653 ... | tri> ¬a ¬b c = Uz
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
654 seq : Uz ≡ supf0 x
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
655 seq with trio< x x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
656 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b refl )
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
657 ... | tri≈ ¬a b ¬c = refl
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
658 ... | tri> ¬a ¬b c = refl
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
659 seq<x : {b : Ordinal } → (b<x : b o< x ) → ZChain1.supf (prev b b<x {y} ay) b ≡ supf0 b
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
660 seq<x {b} b<x with trio< b x
624
d0938f220648 supf again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 613
diff changeset
661 ... | tri< a ¬b ¬c = cong (λ k → ZChain1.supf (prev b k {y} ay) b) o<-irr -- b<x ≡ a
611
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
662 ... | tri≈ ¬a b₁ ¬c = ⊥-elim (¬a b<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 610
diff changeset
663 ... | tri> ¬a ¬b c = ⊥-elim (¬a b<x )
625
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
664 u-mono : {z : Ordinal} {y : Ordinal} → z o≤ y → y o≤ x → supf0 z ⊆' supf0 y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
665 u-mono {z} {y} z≤y y≤x with trio< z x | trio< y x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
666 ... | tri< a ¬b ¬c | t = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
667 ... | tri≈ ¬a b ¬c | t = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
668 ... | tri> ¬a ¬b c | t = ⊥-elim ( osuc-< y≤x {!!} )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
669 u-total : {z : Ordinal} → z o≤ x → IsTotalOrderSet (supf0 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
670 u-total {z} z<x ux uy with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
671 ... | t = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
672 -- with trio< (UZFChain.u ux) (UZFChain.u uy)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
673 -- ... | tri< a ¬b ¬c = ZChain1.f-total (uzc1 uy) {!!} (u-mono (UZFChain.u ux) (UZFChain.u uy)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
674 -- (UZFChain.u<x uy) (ordtrans a <-osuc ) (uzc1 ux) (uzc1 uy) (UZFChain.chain∋z ux)) (UZFChain.chain∋z uy)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
675 -- ... | tri≈ ¬a b ¬c = ZChain1.f-total (uzc1 ux) {!!} (UZFChain.chain∋z ux) (u-mono (UZFChain.u uy) (UZFChain.u ux)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
676 -- (UZFChain.u<x ux) (subst (λ k → k o< osuc (UZFChain.u ux)) b <-osuc) (uzc1 uy) (uzc1 ux) (UZFChain.chain∋z uy))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
677 -- ... | tri> ¬a ¬b c = ZChain1.f-total (uzc1 ux) {!!} (UZFChain.chain∋z ux) (u-mono (UZFChain.u uy) (UZFChain.u ux)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 624
diff changeset
678 -- (UZFChain.u<x ux) (ordtrans c <-osuc) (uzc1 uy) (uzc1 ux) (UZFChain.chain∋z uy))
553
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 552
diff changeset
679
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
680 SZ : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → {y : Ordinal} (ya : odef A y) → ZChain1 A y f (& A)
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
681 SZ f mf = TransFinite {λ z → {y : Ordinal } → (ay : odef A y ) → ZChain1 A y f z } (ind f mf) (& A)
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
682
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
683 zorn00 : Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
684 zorn00 with is-o∅ ( & HasMaximal ) -- we have no Level (suc n) LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
685 ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; A∋maximal = zorn01 ; ¬maximal<x = zorn02 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
686 -- yes we have the maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
687 zorn03 : odef HasMaximal ( & ( ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ) )
606
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
688 zorn03 = ODC.x∋minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) -- Axiom of choice
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
689 zorn01 : A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
690 zorn01 = proj1 zorn03
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
691 zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
692 zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
693 ... | yes ¬Maximal = ⊥-elim ( z04 nmx zorn04) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
694 -- if we have no maximal, make ZChain, which contradict SUP condition
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
695 nmx : ¬ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
696 nmx mx = ∅< {HasMaximal} zc5 ( ≡o∅→=od∅ ¬Maximal ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
697 zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) → odef A y → ¬ (* (& (Maximal.maximal mx)) < * y))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
698 zc5 = ⟪ Maximal.A∋maximal mx , (λ y ay mx<y → Maximal.¬maximal<x mx (subst (λ k → odef A k ) (sym &iso) ay) (subst (λ k → k < * y) *iso mx<y) ) ⟫
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
699 zorn04 : ZChain1 A (& s) (cf nmx) (& A)
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
700 zorn04 = SZ (cf nmx) (cf-is-≤-monotonic nmx) (subst (λ k → odef A k ) &iso as )
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
701
516
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
702 -- usage (see filter.agda )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
703 --
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
704 -- _⊆'_ : ( A B : HOD ) → Set n
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
705 -- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x
482
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 481
diff changeset
706
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
707 -- MaximumSubset : {L P : HOD}
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
708 -- → o∅ o< & L → o∅ o< & P → P ⊆ L
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
709 -- → IsPartialOrderSet P _⊆'_
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
710 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ )
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
711 -- → Maximal P (_⊆'_)
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
712 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆'_} 0<P PO SP