annotate src/zorn.agda @ 532:90f61d55cc54

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 22 Apr 2022 13:56:31 +0900
parents 5ca59261a4aa
children 7325484fc491
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
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
4 import OD
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
5 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) where
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
6
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
7 open import zf
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
8 open import logic
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
9 -- open import partfunc {n} O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
10
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
11 open import Relation.Nullary
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
12 open import Relation.Binary
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
13 open import Data.Empty
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
14 open import Relation.Binary
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
15 open import Relation.Binary.Core
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
16 open import Relation.Binary.PropositionalEquality
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
17 import BAlgbra
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
18
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
19
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
20 open inOrdinal O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
21 open OD O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
22 open OD.OD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
23 open ODAxiom odAxiom
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
24 import OrdUtil
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
25 import ODUtil
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
26 open Ordinals.Ordinals O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
27 open Ordinals.IsOrdinals isOrdinal
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
28 open Ordinals.IsNext isNext
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
29 open OrdUtil O
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
30 open ODUtil O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
31
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
32
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
33 import ODC
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
34
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
35
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
36 open _∧_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
37 open _∨_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
38 open Bool
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
39
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
40
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
41 open HOD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
42
528
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
43 _≤_ : (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
44 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
45
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
46 record Element (A : HOD) : Set (Level.suc n) where
469
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
47 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
48 elm : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
49 is-elm : A ∋ elm
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
50
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
51 open Element
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
52
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
53 _<A_ : {A : HOD} → (x y : Element A ) → Set n
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
54 x <A y = elm x < elm y
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
55 _≡A_ : {A : HOD} → (x y : Element A ) → Set (Level.suc n)
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
56 x ≡A y = elm x ≡ elm y
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
57
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
58 IsPartialOrderSet : ( A : HOD ) → Set (Level.suc n)
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
59 IsPartialOrderSet A = IsStrictPartialOrder (_≡A_ {A}) _<A_
490
00c71d1dc316 IsPartialOrder
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 489
diff changeset
60
492
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
61 open _==_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
62 open _⊆_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
63
495
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
64 isA : { A B : HOD } → B ⊆ A → (x : Element B) → Element A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
65 isA B⊆A x = record { elm = elm x ; is-elm = incl B⊆A (is-elm x) }
494
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 493
diff changeset
66
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
67 ⊆-IsPartialOrderSet : { A B : HOD } → B ⊆ A → IsPartialOrderSet A → IsPartialOrderSet B
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
68 ⊆-IsPartialOrderSet {A} {B} B⊆A PA = record {
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
69 isEquivalence = record { refl = refl ; sym = sym ; trans = trans } ; trans = λ {x} {y} {z} → trans1 {x} {y} {z}
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
70 ; irrefl = λ {x} {y} → irrefl1 {x} {y} ; <-resp-≈ = resp0
493
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 492
diff changeset
71 } where
495
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
72 _<B_ : (x y : Element B ) → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
73 x <B y = elm x < elm y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
74 trans1 : {x y z : Element B} → x <B y → y <B z → x <B z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
75 trans1 {x} {y} {z} x<y y<z = IsStrictPartialOrder.trans PA {isA B⊆A x} {isA B⊆A y} {isA B⊆A z} x<y y<z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
76 irrefl1 : {x y : Element B} → elm x ≡ elm y → ¬ ( x <B y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
77 irrefl1 {x} {y} x=y x<y = IsStrictPartialOrder.irrefl PA {isA B⊆A x} {isA B⊆A y} x=y x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
78 open import Data.Product
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
79 resp0 : ({x y z : Element B} → elm y ≡ elm z → x <B y → x <B z) × ({x y z : Element B} → elm y ≡ elm z → y <B x → z <B x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
80 resp0 = Data.Product._,_ (λ {x} {y} {z} → proj₁ (IsStrictPartialOrder.<-resp-≈ PA) {isA B⊆A x } {isA B⊆A y }{isA B⊆A z })
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
81 (λ {x} {y} {z} → proj₂ (IsStrictPartialOrder.<-resp-≈ PA) {isA B⊆A x } {isA B⊆A y }{isA B⊆A z })
492
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
82
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
83 -- open import Relation.Binary.Properties.Poset as Poset
496
c03d80290855 total of B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 495
diff changeset
84
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
85 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n)
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
86 IsTotalOrderSet A = IsStrictTotalOrder (_≡A_ {A}) _<A_
490
00c71d1dc316 IsPartialOrder
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 489
diff changeset
87
469
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
88 me : { A a : HOD } → A ∋ a → Element A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
89 me {A} {a} lt = record { elm = a ; is-elm = lt }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
90
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
91 A∋x-irr : (A : HOD) {x y : HOD} → x ≡ y → (A ∋ x) ≡ (A ∋ y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
92 A∋x-irr A {x} {y} refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
93
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
94 me-elm-refl : (A : HOD) → (x : Element A) → elm (me {A} (d→∋ A (is-elm x))) ≡ elm x
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
95 me-elm-refl A record { elm = ex ; is-elm = ax } = *iso
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
96
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
97 open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
98
526
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
99 -- Don't use Element other than Order, you'll be in a trouble
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
100 -- postulate -- may be proved by transfinite induction and functional extentionality
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
101 -- ∋x-irr : (A : HOD) {x y : HOD} → x ≡ y → (ax : A ∋ x) (ay : A ∋ y ) → ax ≅ ay
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
102 -- odef-irr : (A : OD) {x y : Ordinal} → x ≡ y → (ax : def A x) (ay : def A y ) → ax ≅ ay
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
103
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
104 -- is-elm-irr : (A : HOD) → {x y : Element A } → elm x ≡ elm y → is-elm x ≅ is-elm y
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
105 -- is-elm-irr A {x} {y} eq = ∋x-irr A eq (is-elm x) (is-elm y )
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
106
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
107 El-irr2 : (A : HOD) → {x y : Element A } → elm x ≡ elm y → is-elm x ≅ is-elm y → x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
108 El-irr2 A {x} {y} refl HE.refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
109
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
110 -- El-irr : {A : HOD} → {x y : Element A } → elm x ≡ elm y → x ≡ y
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
111 -- El-irr {A} {x} {y} eq = El-irr2 A eq ( is-elm-irr A eq )
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
112
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
113 record _Set≈_ (A B : Ordinal ) : Set n where
526
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
114 field
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
115 fun← : {x : Ordinal } → odef (* A) x → Ordinal
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
116 fun→ : {x : Ordinal } → odef (* B) x → Ordinal
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
117 funB : {x : Ordinal } → ( lt : odef (* A) x ) → odef (* B) ( fun← lt )
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
118 funA : {x : Ordinal } → ( lt : odef (* B) x ) → odef (* A) ( fun→ lt )
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
119 fiso← : {x : Ordinal } → ( lt : odef (* B) x ) → fun← ( funA lt ) ≡ x
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
120 fiso→ : {x : Ordinal } → ( lt : odef (* A) x ) → fun→ ( funB lt ) ≡ x
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
121
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
122 open _Set≈_
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
123 record _OS≈_ {A B : HOD} (TA : IsTotalOrderSet A ) (TB : IsTotalOrderSet B ) : Set (Level.suc n) where
526
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
124 field
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
125 iso : (& A ) Set≈ (& B)
526
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
126 fmap : {x y : Ordinal} → (ax : odef A x) → (ay : odef A y) → * x < * y
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
127 → * (fun← iso (subst (λ k → odef k x) (sym *iso) ax)) < * (fun← iso (subst (λ k → odef k y) (sym *iso) ay))
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
128
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
129 Cut< : ( A x : HOD ) → HOD
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
130 Cut< A x = record { od = record { def = λ y → ( odef A y ) ∧ ( x < * y ) } ; odmax = & A
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
131 ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (subst (λ k → odef A k ) (sym &iso) (proj1 lt))) }
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
132
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
133 Cut<T : {A : HOD} → (TA : IsTotalOrderSet A ) ( x : HOD )→ IsTotalOrderSet ( Cut< A x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
134 Cut<T {A} TA x = record { isEquivalence = record { refl = refl ; trans = trans ; sym = sym }
526
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
135 ; trans = λ {x} {y} {z} → IsStrictTotalOrder.trans TA {me (proj1 (is-elm x))} {me (proj1 (is-elm y))} {me (proj1 (is-elm z))} ;
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
136 compare = λ x y → IsStrictTotalOrder.compare TA (me (proj1 (is-elm x))) (me (proj1 (is-elm y))) }
526
7e59e0aeaa73 give up for a while
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
137
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
138 record _OS<_ {A B : HOD} (TA : IsTotalOrderSet A ) (TB : IsTotalOrderSet B ) : Set (Level.suc n) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
139 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
140 x : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
141 iso : TA OS≈ (Cut<T TA x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
142
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
143 -- OS<-cmp : {x : HOD} → Trichotomous {_} {IsTotalOrderSet x} _OS≈_ _OS<_
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
144 -- OS<-cmp A B = {!!}
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
145
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
146 -- tree structure
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
147 data IChain (A : HOD) : Ordinal → Set n where
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
148 ifirst : {ox : Ordinal} → odef A ox → IChain A ox
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
149 inext : {ox oy : Ordinal} → odef A oy → * ox < * oy → IChain A ox → IChain A oy
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
150
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
151 -- * ox < .. < * oy
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
152 ic-connect : {A : HOD} {oy : Ordinal} → (ox : Ordinal) → (iy : IChain A oy) → Set n
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
153 ic-connect {A} ox (ifirst {oy} ay) = Lift n ⊥
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
154 ic-connect {A} ox (inext {oy} {oz} ay y<z iz) = (ox ≡ oy) ∨ ic-connect ox iz
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
155
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
156 ic→odef : {A : HOD} {ox : Ordinal} → IChain A ox → odef A ox
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
157 ic→odef {A} {ox} (ifirst ax) = ax
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
158 ic→odef {A} {ox} (inext ax x<y ic) = ax
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
159
521
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
160 ic→< : {A : HOD} → IsPartialOrderSet A → (x : Ordinal) → odef A x → {y : Ordinal} → (iy : IChain A y) → ic-connect {A} {y} x iy → * x < * y
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
161 ic→< {A} PO x ax {y} (ifirst ay) ()
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
162 ic→< {A} PO x ax {y} (inext ay x<y iy) (case1 refl) = x<y
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
163 ic→< {A} PO x ax {y} (inext {oy} ay x<y iy) (case2 ic) = IsStrictPartialOrder.trans PO
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
164 {me (subst (λ k → odef A k) (sym &iso) ax )} {me (subst (λ k → odef A k) (sym &iso) (ic→odef {A} {oy} iy) ) } {me (subst (λ k → odef A k) (sym &iso) ay) }
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
165 (ic→< {A} PO x ax iy ic ) x<y
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
166
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
167 record IChained (A : HOD) (x y : Ordinal) : Set n where
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
168 field
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
169 iy : IChain A y
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
170 ic : ic-connect x iy
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
171
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
172 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
173 -- all tree from x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
174 --
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
175 IChainSet : (A : HOD) {x : Ordinal} → odef A x → HOD
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
176 IChainSet A {x} ax = record { od = record { def = λ y → odef A y ∧ IChained A x y }
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
177 ; odmax = & A ; <odmax = λ {y} lt → subst (λ k → k o< & A) &iso (c<→o< (subst (λ k → odef A k) (sym &iso) (proj1 lt))) }
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
178
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
179 IChainSet⊆A : {A : HOD} → {x : Ordinal } → (ax : odef A x ) → IChainSet A ax ⊆ A
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
180 IChainSet⊆A {A} x = record { incl = λ {oy} y → proj1 y }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
181
521
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
182 ¬IChained-refl : (A : HOD) {x : Ordinal} → IsPartialOrderSet A → ¬ IChained A x x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
183 ¬IChained-refl A {x} PO record { iy = iy ; ic = ic } = IsStrictPartialOrder.irrefl PO
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
184 {me (subst (λ k → odef A k ) (sym &iso) ic0) } {me (subst (λ k → odef A k ) (sym &iso) ic0) } refl (ic→< {A} PO x ic0 iy ic ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
185 ic0 : odef A x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
186 ic0 = ic→odef {A} iy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
187
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
188 -- there is a y, & y > & x
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
189
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
190 record OSup> (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
191 field
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
192 y : Ordinal
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
193 icy : odef (IChainSet A ax ) y
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
194 y>x : x o< y
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
195
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
196 record IChainSup> (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
197 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
198 y : Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
199 A∋y : odef A y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
200 y>x : * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
201
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
202 -- finite IChain
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
203 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
204 -- tree structured
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
205
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
206 ic→A∋y : (A : HOD) {x y : Ordinal} (ax : A ∋ * x) → odef (IChainSet A ax) y → A ∋ * y
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
207 ic→A∋y A {x} {y} ax ⟪ ay , _ ⟫ = subst (λ k → odef A k) (sym &iso) ay
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
208
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
209 record InfiniteChain (A : HOD) (max : Ordinal) {x : Ordinal} (ax : A ∋ * x) : Set n where
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
210 field
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
211 chain<x : (y : Ordinal ) → odef (IChainSet A ax) y → y o< max
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
212 c-infinite : (y : Ordinal ) → (cy : odef (IChainSet A ax) y )
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
213 → IChainSup> A (ic→A∋y A ax cy)
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
214
528
8facdd7cc65a TransitiveClosure with x <= f x is possible
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 527
diff changeset
215 open import Data.Nat hiding (_<_ ; _≤_ )
510
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
216 import Data.Nat.Properties as NP
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
217 open import nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
218
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
219 data Chain (A : HOD) (s : Ordinal) (next : Ordinal → Ordinal ) : ( x : Ordinal ) → Set n where
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
220 cfirst : odef A s → Chain A s next s
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
221 csuc : (x : Ordinal ) → (ax : odef A x ) → Chain A s next x → odef A (next x) → Chain A s next (next x )
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
222
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
223 ct∈A : (A : HOD ) (s : Ordinal) → (next : Ordinal → Ordinal ) → {x : Ordinal} → Chain A s next x → odef A x
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
224 ct∈A A s next {x} (cfirst x₁) = x₁
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
225 ct∈A A s next {.(next x )} (csuc x ax t anx) = anx
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
226
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
227 --
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
228 -- extract single chain from countable infinite chains
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
229 --
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
230 TransitiveClosure : (A : HOD) (s : Ordinal) → (next : Ordinal → Ordinal ) → HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
231 TransitiveClosure A s next = record { od = record { def = λ x → Chain A s next x } ; odmax = & A ; <odmax = cc01 } where
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
232 cc01 : {y : Ordinal} → Chain A s next y → y o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
233 cc01 {y} cy = subst (λ k → k o< & A ) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso) ( ct∈A A s next cy ) ) )
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
234
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
235 cton0 : (A : HOD ) (s : Ordinal) → (next : Ordinal → Ordinal ) {y : Ordinal } → Chain A s next y → ℕ
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
236 cton0 A s next (cfirst _) = zero
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
237 cton0 A s next (csuc x ax z _) = suc (cton0 A s next z)
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
238 cton : (A : HOD ) (s : Ordinal) → (next : Ordinal → Ordinal ) → Element (TransitiveClosure A s next) → ℕ
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
239 cton A s next y = cton0 A s next (is-elm y)
510
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
240
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
241 cinext : (A : HOD) {x max : Ordinal } → (ax : A ∋ * x ) → (ifc : InfiniteChain A max ax ) → Ordinal → Ordinal
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
242 cinext A ax ifc y with ODC.∋-p O (IChainSet A ax) (* y)
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
243 ... | yes ics-y = IChainSup>.y ( InfiniteChain.c-infinite ifc y (subst (λ k → odef (IChainSet A ax) k) &iso ics-y ))
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
244 ... | no _ = o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
245
522
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 521
diff changeset
246 InFCSet : (A : HOD) → {x max : Ordinal} (ax : A ∋ * x)
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
247 → (ifc : InfiniteChain A max ax ) → HOD
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
248 InFCSet A {x} ax ifc = TransitiveClosure (IChainSet A ax) x (cinext A ax ifc )
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
249
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
250 InFCSet⊆A : (A : HOD) → {x max : Ordinal} (ax : A ∋ * x) → (ifc : InfiniteChain A max ax ) → InFCSet A ax ifc ⊆ A
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
251 InFCSet⊆A A {x} ax ifc = record { incl = λ {y} lt → incl (IChainSet⊆A ax) (
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
252 ct∈A (IChainSet A ax) x (cinext A ax ifc) lt ) }
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
253
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
254 cinext→IChainSup : (A : HOD) {x max : Ordinal } → (ax : A ∋ * x ) → (ifc : InfiniteChain A max ax ) → (y : Ordinal )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
255 → (ay1 : IChainSet A ax ∋ * y ) → IChainSup> A (subst (odef A) (sym &iso) (proj1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1)))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
256 cinext→IChainSup A {x} ax ifc y ay with ODC.∋-p O (IChainSet A ax) (* y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
257 ... | no not = ⊥-elim ( not ay )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
258 ... | yes ay1 = InfiniteChain.c-infinite ifc y (subst (λ k → odef (IChainSet A ax) k) &iso ay )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
259
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
260 TransitiveClosure-is-total : (A : HOD) → {x max : Ordinal} (ax : A ∋ * x)
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
261 → IsPartialOrderSet A
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
262 → (ifc : InfiniteChain A max ax )
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
263 → IsTotalOrderSet ( InFCSet A ax ifc )
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
264 TransitiveClosure-is-total A {x} ax PO ifc = record { isEquivalence = IsStrictPartialOrder.isEquivalence IPO
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
265 ; trans = λ {x} {y} {z} → IsStrictPartialOrder.trans IPO {x} {y} {z} ; compare = cmp } where
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
266 IPO : IsPartialOrderSet (InFCSet A ax ifc )
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
267 IPO = ⊆-IsPartialOrderSet record { incl = λ {y} lt → incl (InFCSet⊆A A {x} ax ifc) lt} PO
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
268 B = IChainSet A ax
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
269 cnext = cinext A ax ifc
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
270 ct02 : {oy : Ordinal} → (y : Chain B x cnext oy ) → A ∋ * oy
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
271 ct02 y = incl (IChainSet⊆A {A} ax) (subst (λ k → odef (IChainSet A ax) k) (sym &iso) (ct∈A B x cnext y) )
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
272 ct-inject : {ox oy : Ordinal} → (x1 : Chain B x cnext ox ) → (y : Chain B x cnext oy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
273 → (cton0 B x cnext x1) ≡ (cton0 B x cnext y) → ox ≡ oy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
274 ct-inject {ox} {ox} (cfirst x) (cfirst x₁) refl = refl
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
275 ct-inject {.(cnext x₀ )} {.(cnext x₃ )} (csuc x₀ ax x₁ x₂) (csuc x₃ ax₁ y x₄) eq = cong cnext ct05 where
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
276 ct06 : {x y : ℕ} → suc x ≡ suc y → x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
277 ct06 refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
278 ct05 : x₀ ≡ x₃
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
279 ct05 = ct-inject x₁ y (ct06 eq)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
280 ct-monotonic : {ox oy : Ordinal} → (x1 : Chain B x cnext ox ) → (y : Chain B x cnext oy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
281 → (cton0 B x cnext x1) Data.Nat.< (cton0 B x cnext y) → * ox < * oy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
282 ct-monotonic {ox} {oy} x1 (csuc oy1 ay y _) (s≤s lt) with NP.<-cmp ( cton0 B x cnext x1 ) ( cton0 B x cnext y )
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
283 ... | tri< a ¬b ¬c = ct07 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
284 ct07 : * ox < * (cnext oy1)
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
285 ct07 with ODC.∋-p O (IChainSet A ax) (* oy1)
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
286 ... | no not = ⊥-elim ( not (subst (λ k → odef (IChainSet A ax) k ) (sym &iso) ay ) )
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
287 ... | yes ay1 = IsStrictPartialOrder.trans PO {me (ct02 x1) } {me (ct02 y)} {me ct031 } (ct-monotonic x1 y a ) ct011 where
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
288 ct031 : A ∋ * (IChainSup>.y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 ) ))
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
289 ct031 = subst (λ k → odef A k ) (sym &iso) (
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
290 IChainSup>.A∋y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
291 ct011 : * oy1 < * ( IChainSup>.y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
292 ct011 = IChainSup>.y>x (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 ))
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
293 ... | tri≈ ¬a b ¬c = ct11 where
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
294 ct11 : * ox < * (cnext oy1)
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
295 ct11 with ODC.∋-p O (IChainSet A ax) (* oy1)
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
296 ... | no not = ⊥-elim ( not (subst (λ k → odef (IChainSet A ax) k ) (sym &iso) ay ) )
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
297 ... | yes ay1 = subst (λ k → * k < _) (sym (ct-inject _ _ b)) ct011 where
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
298 ct011 : * oy1 < * ( IChainSup>.y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
299 ct011 = IChainSup>.y>x (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 ))
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
300 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> lt c )
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
301 ct12 : {y z : Element (TransitiveClosure B x cnext) } → elm y ≡ elm z → elm y < elm z → ⊥
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
302 ct12 {y} {z} y=z y<z = IsStrictPartialOrder.irrefl IPO {y} {z} y=z y<z
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
303 ct13 : {y z : Element (TransitiveClosure B x cnext) } → elm y < elm z → elm z < elm y → ⊥
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
304 ct13 {y} {z} y<z y>z = IsStrictPartialOrder.irrefl IPO {y} {y} refl ( IsStrictPartialOrder.trans IPO {y} {z} {y} y<z y>z )
527
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 526
diff changeset
305 ct17 : (x1 : Element (TransitiveClosure B x cnext)) → Chain B x cnext (& (elm x1))
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
306 ct17 x1 = is-elm x1
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
307 cmp : Trichotomous _ _
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
308 cmp x1 y with NP.<-cmp (cton B x cnext x1) (cton B x cnext y)
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
309 ... | tri< a ¬b ¬c = tri< ct04 ct14 ct15 where
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
310 ct04 : elm x1 < elm y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
311 ct04 = subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a)
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
312 ct14 : ¬ elm x1 ≡ elm y
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
313 ct14 eq = ct12 {x1} {y} eq (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a ) )
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
314 ct15 : ¬ (elm y < elm x1)
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
315 ct15 lt = ct13 {y} {x1} lt (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a ) )
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
316 ... | tri≈ ¬a b ¬c = tri≈ (ct12 {x1} {y} ct16) ct16 (ct12 {y} {x1} (sym ct16)) where
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
317 ct16 : elm x1 ≡ elm y
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
318 ct16 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) (ct-inject {& (elm x1)} {& (elm y)} (is-elm x1) (is-elm y) b ))
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
319 ... | tri> ¬a ¬b c = tri> ct15 ct14 ct04 where
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
320 ct04 : elm y < elm x1
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
321 ct04 = subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c)
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
322 ct14 : ¬ elm x1 ≡ elm y
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
323 ct14 eq = ct12 {y} {x1} (sym eq) (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c ) )
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
324 ct15 : ¬ (elm x1 < elm y)
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
325 ct15 lt = ct13 {x1} {y} lt (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c ) )
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
326
502
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
327 record IsFC (A : HOD) {x : Ordinal} (ax : A ∋ * x) (y : Ordinal) : Set n where
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
328 field
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
329 icy : odef (IChainSet A ax) y
520
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 519
diff changeset
330 c-finite : ¬ IChainSup> A (subst (λ k → odef A k ) (sym &iso) (proj1 icy) )
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
331
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
332 record Maximal ( A : HOD ) : Set (Level.suc n) where
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
333 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
334 maximal : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
335 A∋maximal : A ∋ maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
336 ¬maximal<x : {x : HOD} → A ∋ x → ¬ maximal < x -- A is Partial, use negative
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
337
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
338 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
339 -- possible three cases in a limit ordinal step
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
340 --
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
341 -- case 1) < goes x o<
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
342 -- case 2) no > x in some chain ( maximal )
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
343 -- case 3) countably infinite chain below x
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
344 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
345 Zorn-lemma-3case : { A : HOD }
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
346 → o∅ o< & A
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
347 → IsPartialOrderSet A
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
348 → (x : Ordinal ) → (ax : odef A x) → OSup> A (d→∋ A ax) ∨ Maximal A ∨ InfiniteChain A x (d→∋ A ax)
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
349 Zorn-lemma-3case {A} 0<A PO x ax = zc2 where
499
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 498
diff changeset
350 Gtx : HOD
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
351 Gtx = record { od = record { def = λ y → odef ( IChainSet A ax ) y ∧ ( x o< y ) } ; odmax = & A
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
352 ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (d→∋ A (proj1 (proj1 lt)))) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
353 HG : HOD
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
354 HG = record { od = record { def = λ y → odef A y ∧ IsFC A (d→∋ A ax ) y } ; odmax = & A
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
355 ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (d→∋ A (proj1 lt) )) }
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
356 zc2 : OSup> A (d→∋ A ax) ∨ Maximal A ∨ InfiniteChain A x (d→∋ A ax )
499
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 498
diff changeset
357 zc2 with is-o∅ (& Gtx)
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
358 ... | no not = case1 record { y = & y ; icy = zc4 ; y>x = proj2 zc3 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
359 y : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
360 y = ODC.minimal O Gtx (λ eq → not (=od∅→≡o∅ eq))
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
361 zc3 : odef ( IChainSet A ax ) (& y) ∧ ( x o< (& y ))
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
362 zc3 = ODC.x∋minimal O Gtx (λ eq → not (=od∅→≡o∅ eq))
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
363 zc4 : odef (IChainSet A (subst (OD.def (od A)) (sym &iso) ax)) (& y)
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
364 zc4 = ⟪ proj1 (proj1 zc3) , (subst (λ k → IChained A k (& y)) (sym &iso) (proj2 (proj1 zc3))) ⟫
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
365 ... | yes nogt with is-o∅ (& HG)
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
366 ... | no finite-chain = case2 (case1 record { maximal = y ; A∋maximal = proj1 zc3 ; ¬maximal<x = zc4 } ) where
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
367 y : HOD
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
368 y = ODC.minimal O HG (λ eq → finite-chain (=od∅→≡o∅ eq))
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
369 zc3 : odef A (& y) ∧ IsFC A (d→∋ A ax ) (& y)
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
370 zc3 = ODC.x∋minimal O HG (λ eq → finite-chain (=od∅→≡o∅ eq))
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
371 zc4 : {z : HOD} → A ∋ z → ¬ (y < z)
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
372 zc4 {z} az y<z = IsFC.c-finite (proj2 zc3) record { y = & z ; A∋y = az ; y>x = subst₂ (λ j k → j < k ) (sym *iso) (sym *iso) y<z }
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
373 ... | yes inifite = case2 (case2 record { c-infinite = zc91 ; chain<x = zc10 } ) where
518
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 517
diff changeset
374 B : HOD
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
375 B = IChainSet A ax -- (me (subst (OD.def (od A)) (sym &iso) (is-elm x)))
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
376 B1 : HOD
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
377 B1 = IChainSet A (subst (OD.def (od A)) (sym &iso) ax)
518
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 517
diff changeset
378 Nx : (y : Ordinal) → odef A y → HOD
520
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 519
diff changeset
379 Nx y ay = record { od = record { def = λ x → odef A x ∧ ( * y < * x ) } ; odmax = & A
518
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 517
diff changeset
380 ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (d→∋ A (proj1 lt))) }
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
381 zc10 : (y : Ordinal) → odef (IChainSet A (subst (OD.def (od A)) (sym &iso) ax)) y → y o< x
521
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
382 zc10 oy icsy = zc21 where
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
383 zc20 : (y : HOD) → (IChainSet A ax) ∋ y → x o< & y → ⊥
521
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
384 zc20 y icsy lt = ¬A∋x→A≡od∅ Gtx ⟪ icsy , lt ⟫ nogt
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
385 zc22 : IChainSet A ax ∋ * oy
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
386 zc22 = ⟪ subst (λ k → odef A k) (sym &iso) (proj1 icsy) , subst₂ (λ j k → IChained A j k ) &iso (sym &iso) (proj2 icsy) ⟫
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
387 zc21 : oy o< x
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
388 zc21 with trio< oy x
521
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 520
diff changeset
389 ... | tri< a ¬b ¬c = a
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
390 ... | tri≈ ¬a b ¬c = ⊥-elim (¬IChained-refl A PO (subst₂ (λ j k → IChained A j k ) &iso b (proj2 icsy)) )
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
391 ... | tri> ¬a ¬b c = ⊥-elim ( zc20 (* oy) zc22 (subst (λ k → x o< k) (sym &iso) c ))
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
392 zc91 : (y : Ordinal) (cy : odef B1 y) → IChainSup> A (ic→A∋y A (subst (OD.def (od A)) (sym &iso) ax) cy)
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
393 zc91 y cy with is-o∅ (& (Nx y (proj1 cy) ))
519
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 518
diff changeset
394 ... | yes no-next = ⊥-elim zc16 where
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
395 zc18 : ¬ IChainSup> A (subst (odef A) (sym &iso) (proj1 (subst (λ k → odef (IChainSet A (d→∋ A ax)) k) (sym &iso) cy)))
520
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 519
diff changeset
396 zc18 ics = ¬A∋x→A≡od∅ (Nx y (proj1 cy) ) ⟪ subst (λ k → odef A k ) (sym &iso) (IChainSup>.A∋y ics)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 519
diff changeset
397 , subst₂ (λ j k → j < k ) *iso (cong (*) (sym &iso))( IChainSup>.y>x ics) ⟫ no-next
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
398 zc17 : IsFC A {x} (d→∋ A ax) (& (* y))
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
399 zc17 = record { icy = subst (λ k → odef (IChainSet A (d→∋ A ax)) k) (sym &iso) cy ; c-finite = zc18 }
519
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 518
diff changeset
400 zc16 : ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 518
diff changeset
401 zc16 = ¬A∋x→A≡od∅ HG ⟪ subst (λ k → odef A k ) (sym &iso) (proj1 cy ) , zc17 ⟫ inifite
520
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 519
diff changeset
402 ... | no not = record { y = & zc13 ; A∋y = proj1 zc12 ; y>x = proj2 zc12 } where
519
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 518
diff changeset
403 zc13 = ODC.minimal O (Nx y (proj1 cy)) (λ eq → not (=od∅→≡o∅ eq ))
520
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 519
diff changeset
404 zc12 : odef A (& zc13 ) ∧ ( * y < * ( & zc13 ))
519
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 518
diff changeset
405 zc12 = ODC.x∋minimal O (Nx y (proj1 cy)) (λ eq → not (=od∅→≡o∅ eq ))
499
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 498
diff changeset
406
517
7b99c8944df7 chain total complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 516
diff changeset
407 all-climb-case : { A : HOD } → (0<A : o∅ o< & A) → IsPartialOrderSet A
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
408 → (( x : Ordinal ) → (ax : odef A (& (* x))) → OSup> A ax )
524
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 523
diff changeset
409 → (x : HOD) ( ax : A ∋ x )
525
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
410 → InfiniteChain A (& A) (d→∋ A ax)
524
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 523
diff changeset
411 all-climb-case {A} 0<A PO climb x ax = record { c-infinite = ac00 ; chain<x = ac01 } where
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
412 B = IChainSet A ax
524
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 523
diff changeset
413 ac01 : (y : Ordinal) → odef (IChainSet A (d→∋ A ax)) y → y o< & A
522
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 521
diff changeset
414 ac01 y ⟪ ay , _ ⟫ = subst (λ k → k o< & A ) &iso (c<→o< (subst (λ k → odef A k ) (sym &iso) ay) )
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
415 ac00 : (y : Ordinal) (cy : odef (IChainSet A (d→∋ A ax)) y) → IChainSup> A (ic→A∋y A (d→∋ A ax) cy)
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
416 ac00 y cy = record { y = z ; A∋y = az ; y>x = y<z} where
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
417 ay : odef A (& (* y))
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
418 ay = subst (λ k → odef A k) (sym &iso) (proj1 cy)
522
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 521
diff changeset
419 z : Ordinal
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
420 z = OSup>.y ( climb y ay)
522
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 521
diff changeset
421 az : odef A z
523
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
422 az = subst (λ k → odef A k) &iso ( incl (IChainSet⊆A {A} ay ) (subst (λ k → odef (IChainSet A ay) k ) (sym &iso) (OSup>.icy ( climb y ay))))
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
423 icy : odef (IChainSet A ay ) z
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
424 icy = OSup>.icy ( climb y ay )
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
425 y<z : * y < * z
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
426 y<z = ic→< {A} PO y (subst (λ k → odef A k) &iso ay) (IChained.iy (proj2 icy))
f351c183e712 all-climb-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
427 (subst (λ k → ic-connect k (IChained.iy (proj2 icy))) &iso (IChained.ic (proj2 icy)))
522
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 521
diff changeset
428
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
429 -- <-TransFinite : ( A : HOD ) → IsTotalOrderSet A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
430 -- → ( (x : Ordinal) → ((y : Ordinal) → y o< x → ZChain A y) → ZChain A x ) → (x : Ordinal ) → ZChain A x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
431 -- <-TransFinite A TA ind x = TransFinite {ZChain A} ind x
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
432
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
433 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
434 -- inductive maxmum tree from x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
435 -- tree structure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
436 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
437
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
438 ≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set (Level.suc n)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
439 ≤-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
440
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
441 record ZChain ( A : HOD ) {x : Ordinal} (ax : A ∋ * x) (z : Ordinal) : Set (Level.suc n) where
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
442 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
443 chain : HOD
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
444 chain⊆A : chain ⊆ A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
445 f-total : ( f : Ordinal → Ordinal ) → ≤-monotonic-f A f → IsTotalOrderSet chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
446 is-max : ( f : Ordinal → Ordinal ) → ≤-monotonic-f A f → {a b : Ordinal } → odef chain a → a o< z → * a < * b → odef chain b
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
447
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
448 record SUP ( A B : HOD ) : Set (Level.suc n) where
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
449 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
450 sup : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
451 A∋maximal : A ∋ sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
452 x<sup : {x : HOD} → B ∋ x → (x ≡ sup ) ∨ (x < sup ) -- B is Total, use positive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
453
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
454 Zorn-lemma : { A : HOD }
464
5acf6483a9e3 Zorn lemma start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 447
diff changeset
455 → o∅ o< & A
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
456 → IsPartialOrderSet A
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
457 → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
458 → Maximal A
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
459 Zorn-lemma {A} 0<A PO supP = zorn00 where
493
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 492
diff changeset
460 z01 : {a b : HOD} → A ∋ a → A ∋ b → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
496
c03d80290855 total of B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 495
diff changeset
461 z01 {a} {b} A∋a A∋b (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO {me A∋b} {me A∋a} (sym a=b) b<a
524
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 523
diff changeset
462 z01 {a} {b} A∋a A∋b (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO {me A∋b} {me A∋b} refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 523
diff changeset
463 (IsStrictPartialOrder.trans PO {me A∋b} {me A∋a} {me A∋b} b<a a<b)
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
464 s : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
465 s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
466 sa : A ∋ * ( & s )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
467 sa = 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 )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
468 HasMaximal : HOD
531
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
469 HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) → odef A m → ¬ (* x < * m)) } ; odmax = & A ; <odmax = z07 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
470 z07 : {y : Ordinal} → odef A y ∧ ((m : Ordinal) → odef A m → ¬ (* y < * m)) → y o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
471 z07 {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
472 no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → ((m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m) )) → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
473 no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ? )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
474 Gtx : { x : HOD} → A ∋ x → HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
475 Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = {!!} }
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
476 cf : ¬ Maximal A → Ordinal → Ordinal
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
477 cf nmx x with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
478 ... | no _ = o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
479 ... | yes ax with is-o∅ (& ( Gtx ax ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
480 ... | yes nogt = ⊥-elim (no-maximum ? x x-is-maximal ) where -- no larger element, so it is maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
481 x-is-maximal : (m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
482 x-is-maximal m am = ⟪ subst (λ k → odef A k) &iso ax , ¬x<m ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
483 ¬x<m : ¬ (* x < * m)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
484 ¬x<m x<m = ∅< {Gtx ax} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫ (≡o∅→=od∅ nogt)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
485 ... | no not = & (ODC.minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq)))
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
486 cf-is-<-monotonic : (nmx : ¬ Maximal A ) → (x : Ordinal) → odef A x → ( * x < * (cf nmx x) ) ∧ odef A (cf nmx x )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
487 cf-is-<-monotonic nmx x ax = ⟪ {!!} , {!!} ⟫
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
488 cf-is-≤-monotonic : (nmx : ¬ Maximal A ) → ≤-monotonic-f A ( cf nmx )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
489 cf-is-≤-monotonic nmx x ax = ⟪ case2 (proj1 ( cf-is-<-monotonic nmx x ax )) , proj2 ( cf-is-<-monotonic nmx x ax ) ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
490 zsup : (zc : ZChain A sa (& A)) → ( f : Ordinal → Ordinal ) → ≤-monotonic-f A f → SUP A (ZChain.chain zc)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
491 zsup zc f mf = supP (ZChain.chain zc) (ZChain.chain⊆A zc) ( ZChain.f-total zc f mf )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
492 -- zsup zc f mf = & ( SUP.sup (supP (ZChain.chain zc) (ZChain.chain⊆A zc) ( ZChain.f-total zc f mf ) ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
493 A∋zsup : (zc : ZChain A sa (& A)) → (nmx : ¬ Maximal A ) → A ∋ * ( & ( SUP.sup (zsup zc (cf nmx) (cf-is-≤-monotonic nmx)) ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
494 A∋zsup zc nmx = subst (λ k → odef A (& k )) (sym *iso) ( SUP.A∋maximal (zsup zc (cf nmx) (cf-is-≤-monotonic nmx)) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
495 z03 : (zc : ZChain A sa (& A)) → ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → f (& ( SUP.sup (zsup zc f mf ))) ≡ & (SUP.sup (zsup zc f mf ))
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
496 z03 = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
497 z04 : (zc : ZChain A sa (& A)) → ¬ Maximal A → ⊥
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
498 z04 zc nmx = z01 {* (cf nmx c)} {* c} {!!} (A∋zsup zc nmx ) (case1 ( cong (*)( z03 zc (cf nmx) (cf-is-≤-monotonic nmx ))))
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
499 (proj1 (cf-is-<-monotonic nmx c ((subst λ k → odef A k ) &iso (A∋zsup zc nmx )))) where
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
500 c = & (SUP.sup (zsup zc (cf nmx) (cf-is-≤-monotonic nmx)))
478
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 477
diff changeset
501 -- ZChain is not compatible with the SUP condition
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
502 ind : (x : Ordinal) → ((y : Ordinal) → y o< x → ZChain A sa y ∨ Maximal A )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
503 → ZChain A sa x ∨ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
504 ind x prev with Oprev-p x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
505 ... | yes op with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
506 ... | no ¬Ax = zc1 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
507 -- we have previous ordinal and ¬ A ∋ x, use previous Zchain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
508 px = Oprev.oprev op
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
509 zc1 : ZChain A sa x ∨ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
510 zc1 with prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
511 ... | case2 x = case2 x -- we have the Maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
512 ... | case1 zc = case1 {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
513 ... | yes ax = zc4 where -- we have previous ordinal and A ∋ x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
514 px = Oprev.oprev op
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
515 zc1 : OSup> A (subst (OD.def (od A)) (sym &iso) (subst (OD.def (od A)) &iso ax)) → ZChain A sa x ∨ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
516 zc1 os with prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
517 ... | case2 mx = case2 mx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
518 ... | case1 zc = case1 {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
519 zc4 : ZChain A sa x ∨ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
520 zc4 with Zorn-lemma-3case 0<A PO x (subst (λ k → odef A k) &iso ax )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
521 ... | case1 y>x = zc1 y>x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
522 ... | case2 (case1 mx) = case2 mx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
523 ... | case2 (case2 ic) with prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
524 ... | case2 mx = case2 mx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
525 ... | case1 zc = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
526 ind x prev | no ¬ox with trio< (& A) x --- limit ordinal case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
527 ... | t = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
528
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
529 zorn00 : Maximal A
531
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
530 zorn00 with is-o∅ ( & HasMaximal ) -- we have no Level (suc n) LEM
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
531 ... | 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: 529
diff changeset
532 -- yes we have the maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
533 zorn03 : odef HasMaximal ( & ( ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
534 zorn03 = ODC.x∋minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
535 zorn01 : A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
531
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
536 zorn01 = proj1 zorn03
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
537 zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x)
531
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
538 zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x )
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
539 ... | yes ¬Maximal = ⊥-elim ( z04 zorn03 zorn04 ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
540 -- if we have no maximal, make ZChain, which contradict SUP condition
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
541 zorn04 : ¬ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
542 zorn04 mx = ∅< {HasMaximal} zc5 ( ≡o∅→=od∅ ¬Maximal ) where
531
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
543 zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) → odef A y → ¬ (* (& (Maximal.maximal mx)) < * y))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 530
diff changeset
544 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) ) ⟫
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
545 zorn03 : ZChain A sa (& A)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
546 zorn03 with TransFinite ind (& A)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
547 ... | case1 zc = zc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
548 ... | case2 mx = ⊥-elim ( zorn04 mx )
464
5acf6483a9e3 Zorn lemma start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 447
diff changeset
549
516
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
550 -- usage (see filter.agda )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
551 --
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
552 -- _⊆'_ : ( A B : HOD ) → Set n
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
553 -- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x
482
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 481
diff changeset
554
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
555 -- MaximumSubset : {L P : HOD}
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
556 -- → o∅ o< & L → o∅ o< & P → P ⊆ L
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
557 -- → IsPartialOrderSet P _⊆'_
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
558 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ )
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
559 -- → Maximal P (_⊆'_)
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
560 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆'_} 0<P PO SP