annotate src/zorn.agda @ 744:d92ad9e365b6

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