annotate src/zorn.agda @ 515:5faeae7cfe22

ε-induction does not work on Zorn
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 16 Apr 2022 12:10:09 +0900
parents 97c8abf28706
children 286016848403
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
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
43 record Element (A : HOD) : Set (Level.suc n) where
469
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
44 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
45 elm : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
46 is-elm : A ∋ elm
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
47
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
48 open Element
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
49
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
50 _<A_ : {A : HOD} → (x y : Element A ) → Set n
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
51 x <A y = elm x < elm y
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
52 _≡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
53 x ≡A y = elm x ≡ elm y
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
54
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
55 IsPartialOrderSet : ( A : HOD ) → Set (Level.suc n)
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
56 IsPartialOrderSet A = IsStrictPartialOrder (_≡A_ {A}) _<A_
490
00c71d1dc316 IsPartialOrder
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 489
diff changeset
57
492
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
58 open _==_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
59 open _⊆_
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
60
495
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
61 isA : { A B : HOD } → B ⊆ A → (x : Element B) → Element A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
62 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
63
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
64 ⊆-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
65 ⊆-IsPartialOrderSet {A} {B} B⊆A PA = record {
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
66 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
67 ; irrefl = λ {x} {y} → irrefl1 {x} {y} ; <-resp-≈ = resp0
493
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 492
diff changeset
68 } where
495
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
69 _<B_ : (x y : Element B ) → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
70 x <B y = elm x < elm y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
71 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
72 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
73 irrefl1 : {x y : Element B} → elm x ≡ elm y → ¬ ( x <B y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
74 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
75 open import Data.Product
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 494
diff changeset
76 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
77 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
78 (λ {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
79
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
80 -- 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
81
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
82 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n)
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
83 IsTotalOrderSet A = IsStrictTotalOrder (_≡A_ {A}) _<A_
490
00c71d1dc316 IsPartialOrder
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 489
diff changeset
84
469
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
85 me : { A a : HOD } → A ∋ a → Element A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
86 me {A} {a} lt = record { elm = a ; is-elm = lt }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 468
diff changeset
87
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
88 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
89 A∋x-irr A {x} {y} refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
90
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
91 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
92 me-elm-refl A record { elm = ex ; is-elm = ax } = *iso
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
93
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
94 open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
95
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
96 postulate
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
97 ∋x-irr : (A : HOD) {x y : HOD} → x ≡ y → (ax : A ∋ x) (ay : A ∋ y ) → ax ≅ ay
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
98 odef-irr : (A : OD) {x y : Ordinal} → x ≡ y → (ax : def A x) (ay : def A y ) → ax ≅ ay
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
99
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
100 is-elm-irr : (A : HOD) → {x y : Element A } → elm x ≡ elm y → is-elm x ≅ is-elm y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
101 is-elm-irr A {x} {y} eq = ∋x-irr A eq (is-elm x) (is-elm y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
102
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
103 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
104 El-irr2 A {x} {y} refl HE.refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
105
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
106 El-irr : {A : HOD} → {x y : Element A } → elm x ≡ elm y → x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
107 El-irr {A} {x} {y} eq = El-irr2 A eq ( is-elm-irr A eq )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
108
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
109 record ZChain ( A : HOD ) (y : Ordinal) : Set (Level.suc n) where
464
5acf6483a9e3 Zorn lemma start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 447
diff changeset
110 field
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
111 max : HOD
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
112 A∋max : A ∋ max
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
113 y<max : y o< & max
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
114 chain : HOD
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
115 chain⊆A : chain ⊆ A
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
116 total : IsTotalOrderSet chain
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
117 chain-max : (x : HOD) → chain ∋ x → (x ≡ max ) ∨ ( x < max )
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
118
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
119 data IChain (A : HOD) : Ordinal → Set n where
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
120 ifirst : {ox : Ordinal} → odef A ox → IChain A ox
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
121 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
122
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
123 -- * ox < .. < * oy
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
124 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
125 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
126 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
127
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
128 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
129 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
130 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
131
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
132 ic→< : {A : HOD} → (IsPartialOrderSet A) → (x : Ordinal) → odef A x → {y : Ordinal} → (iy : IChain A y) → ic-connect {A} {y} x iy → * x < * y
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
133 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
134 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
135 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
136 {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
137 (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
138
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
139 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
140 field
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
141 iy : IChain A y
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
142 ic : ic-connect x iy
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
143
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
144 IChainSet : {A : HOD} → Element A → HOD
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
145 IChainSet {A} ax = record { od = record { def = λ y → odef A y ∧ IChained A (& (elm ax)) y }
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
146 ; 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
147
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
148 IChainSet⊆A : {A : HOD} → (x : Element A ) → IChainSet x ⊆ A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
149 IChainSet⊆A {A} x = record { incl = λ {oy} y → proj1 y }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
150
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
151 -- there is a y, & y > & x
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
152
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
153 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
154 field
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
155 y : Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
156 icy : odef (IChainSet {A} (me ax)) y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
157 y>x : x o< y
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
158
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
159 record IChainSup> (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
160 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
161 y : Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
162 A∋y : odef A y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
163 y>x : * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
164
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
165 -- finite IChain
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
166
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
167 ic→A∋y : (A : HOD) {x y : Ordinal} (ax : A ∋ * x) → odef (IChainSet {A} (me ax)) y → A ∋ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
168 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
169
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
170 record InFiniteIChain (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
171 field
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
172 chain<x : (y : Ordinal ) → odef (IChainSet {A} (me ax)) y → y o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
173 c-infinite : (y : Ordinal ) → (cy : odef (IChainSet {A} (me ax)) y )
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
174 → IChainSup> A (ic→A∋y A ax cy)
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
175
510
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
176 open import Data.Nat hiding (_<_)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
177 import Data.Nat.Properties as NP
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
178 open import nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
179
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
180 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
181 cfirst : odef A s → Chain A s next s
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
182 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
183
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
184 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
185 ct∈A A s next {x} (cfirst x₁) = x₁
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
186 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
187
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
188 ChainClosure : (A : HOD) (s : Ordinal) → (next : Ordinal → Ordinal ) → HOD
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
189 ChainClosure A s next = record { od = record { def = λ x → Chain A s next x } ; odmax = & A ; <odmax = cc01 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
190 cc01 : {y : Ordinal} → Chain A s next y → y o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
191 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
192
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
193 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
194 cton0 A s next (cfirst _) = zero
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
195 cton0 A s next (csuc x ax z _) = suc (cton0 A s next z)
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
196 cton : (A : HOD ) (s : Ordinal) → (next : Ordinal → Ordinal ) → Element (ChainClosure A s next) → ℕ
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
197 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
198
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
199 cinext : (A : HOD) {x : Ordinal } → (ax : A ∋ * x ) → (ifc : InFiniteIChain A ax ) → Ordinal → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
200 cinext A ax ifc y with ODC.∋-p O (IChainSet (me ax)) (* y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
201 ... | yes ics-y = IChainSup>.y ( InFiniteIChain.c-infinite ifc y (subst (λ k → odef (IChainSet (me ax)) k) &iso ics-y ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
202 ... | no _ = o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
203
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
204 InFCSet : (A : HOD) → {x : Ordinal} (ax : A ∋ * x)
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
205 → (ifc : InFiniteIChain A ax ) → HOD
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
206 InFCSet A {x} ax ifc = ChainClosure (IChainSet {A} (me ax)) x {!!} -- (λ y → IChainSup>.y ( InFiniteIChain.c-infinite ifc y ay ) )
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
207
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
208 InFCSet⊆A : (A : HOD) → {x : Ordinal} (ax : A ∋ * x) → (ifc : InFiniteIChain A ax ) → InFCSet A ax ifc ⊆ A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
209 InFCSet⊆A A {x} ax ifc = record { incl = λ {y} lt → incl (IChainSet⊆A (me ax)) (
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
210 ct∈A (IChainSet {A} (me ax)) x {!!} lt ) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
211 -- ct∈A (IChainSet {A} (me ax)) x (λ y ay → IChainSup>.y ( InFiniteIChain.c-infinite ifc y ay ) ) lt ) }
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
212
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
213 ChainClosure-is-total : (A : HOD) → {x : Ordinal} (ax : A ∋ * x)
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
214 → IsPartialOrderSet A
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
215 → (ifc : InFiniteIChain A ax )
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
216 → IsTotalOrderSet ( InFCSet A ax ifc )
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
217 ChainClosure-is-total A {x} ax PO ifc = record { isEquivalence = IsStrictPartialOrder.isEquivalence IPO
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
218 ; 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
219 IPO : IsPartialOrderSet (InFCSet A ax ifc )
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
220 IPO = ⊆-IsPartialOrderSet record { incl = λ {y} lt → incl (InFCSet⊆A A {x} ax ifc) lt} PO
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
221 B = IChainSet {A} (me ax)
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
222 cnext = cinext A ax ifc
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
223 ct02 : {oy : Ordinal} → (y : Chain B x cnext oy ) → A ∋ * oy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
224 ct02 y = incl (IChainSet⊆A {A} (me ax)) (subst (λ k → odef (IChainSet (me ax)) k) (sym &iso) (ct∈A B x cnext y) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
225 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
226 → (cton0 B x cnext x1) ≡ (cton0 B x cnext y) → ox ≡ oy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
227 ct-inject {ox} {ox} (cfirst x) (cfirst x₁) refl = refl
514
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
228 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
229 ct06 : {x y : ℕ} → suc x ≡ suc y → x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
230 ct06 refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
231 ct05 : x₀ ≡ x₃
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
232 ct05 = ct-inject x₁ y (ct06 eq)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
233 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
234 → (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
235 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
236 ... | tri< a ¬b ¬c = ct07 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
237 ct07 : * ox < * (cnext oy1)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
238 ct07 with ODC.∋-p O (IChainSet {A} (me ax)) (* oy1)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
239 ... | no not = ⊥-elim ( not (subst (λ k → odef (IChainSet {A} (me ax)) k ) (sym &iso) ay ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
240 ... | yes ay1 = IsStrictPartialOrder.trans PO {me (ct02 x1) } {me (ct02 y)} {me ct031 } (ct-monotonic x1 y a ) ct011 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
241 ct031 : A ∋ * (IChainSup>.y (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 ) ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
242 ct031 = subst (λ k → odef A k ) (sym &iso) (
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
243 IChainSup>.A∋y (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
244 ct011 : * oy1 < * ( IChainSup>.y (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 513
diff changeset
245 ct011 = IChainSup>.y>x (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 ))
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
246 ... | tri≈ ¬a b ¬c = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
247 ... | tri> ¬a ¬b c = {!!}
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
248 cmp : Trichotomous _ _
513
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
249 cmp x1 y with NP.<-cmp (cton B x cnext x1) (cton B x cnext y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
250 ... | tri< a ¬b ¬c = tri< ct04 {!!} {!!} where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
251 ct04 : elm x1 < elm y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 512
diff changeset
252 ct04 = subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a)
510
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
253 ... | tri≈ ¬a b ¬c = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
254 ... | tri> ¬a ¬b c = {!!}
509
72ea26339f66 chain closure
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
255
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
256
502
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
257 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
258 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
259 icy : odef (IChainSet {A} (me ax)) y
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
260 c-finite : ¬ IChainSup> A ax
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
261
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
262 record Maximal ( A : HOD ) : Set (Level.suc n) where
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
263 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
264 maximal : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
265 A∋maximal : A ∋ maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
266 ¬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
267
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
268 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
269 -- possible three cases in a limit ordinal step
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
270 --
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
271 -- case 1) < goes > x (will contradic in the transfinite induction )
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
272 -- case 2) no > x in some chain ( maximal )
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
273 -- case 3) countably infinite chain below x (will be prohibited by sup condtion )
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
274 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
275 Zorn-lemma-3case : { A : HOD }
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
276 → o∅ o< & A
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
277 → IsPartialOrderSet A
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
278 → (x : Element A) → OSup> A (d→∋ A (is-elm x)) ∨ Maximal A ∨ InFiniteIChain A (d→∋ A (is-elm x))
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
279 Zorn-lemma-3case {A} 0<A PO x = zc2 where
499
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 498
diff changeset
280 Gtx : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 498
diff changeset
281 Gtx = record { od = record { def = λ y → odef ( IChainSet x ) y ∧ ( & (elm x) o< y ) } ; odmax = & A
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
282 ; <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
283 HG : HOD
502
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
284 HG = record { od = record { def = λ y → odef A y ∧ IsFC A (d→∋ A (is-elm x) ) y } ; odmax = & A
501
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
285 ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (d→∋ A (proj1 lt) )) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
286 zc2 : OSup> A (d→∋ A (is-elm x)) ∨ Maximal A ∨ InFiniteIChain A (d→∋ A (is-elm x))
499
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 498
diff changeset
287 zc2 with is-o∅ (& Gtx)
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
288 ... | 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
289 y : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
290 y = ODC.minimal O Gtx (λ eq → not (=od∅→≡o∅ eq))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
291 zc3 : odef ( IChainSet x ) (& y) ∧ ( & (elm x) o< (& y ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
292 zc3 = ODC.x∋minimal O Gtx (λ eq → not (=od∅→≡o∅ eq))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
293 zc4 : odef (IChainSet (me (subst (OD.def (od A)) (sym &iso) (is-elm x)))) (& y)
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
294 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
295 ... | yes nogt with is-o∅ (& HG)
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
296 ... | 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
297 y : HOD
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
298 y = ODC.minimal O HG (λ eq → finite-chain (=od∅→≡o∅ eq))
504
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
299 zc3 : odef A (& y) ∧ IsFC A (d→∋ A (is-elm x) ) (& y)
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
300 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
301 zc5 : odef (IChainSet {A} (me (d→∋ A (is-elm x) ))) (& y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
302 zc5 = IsFC.icy (proj2 zc3)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
303 zc4 : {z : HOD} → A ∋ z → ¬ (y < z)
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
304 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) zc6 } where
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
305 zc8 : ic-connect (& (* (& (elm x)))) (IChained.iy (proj2 zc5))
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
306 zc8 = IChained.ic (proj2 zc5)
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
307 zc7 : elm x < y
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
308 zc7 = subst₂ (λ j k → j < k ) *iso *iso ( ic→< {A} PO (& (elm x)) (is-elm x) (IChained.iy (proj2 zc5))
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
309 (subst (λ k → ic-connect (& k) (IChained.iy (proj2 zc5)) ) (me-elm-refl A x) (IChained.ic (proj2 zc5)) ) )
505
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 504
diff changeset
310 zc6 : elm x < z
506
dfcb98151273 2 cases in 3 cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 505
diff changeset
311 zc6 = IsStrictPartialOrder.trans PO {x} {me (proj1 zc3)} {me az} zc7 y<z
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
312 ... | yes inifite = case2 (case2 record { chain<x = {!!} ; c-infinite = {!!} } )
499
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 498
diff changeset
313
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
314
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
315 record SUP ( A B : HOD ) : Set (Level.suc n) where
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
316 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
317 sup : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
318 A∋maximal : A ∋ sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
319 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
320
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
321 Zorn-lemma : { A : HOD }
464
5acf6483a9e3 Zorn lemma start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 447
diff changeset
322 → o∅ o< & A
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
323 → IsPartialOrderSet A
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
324 → ( ( 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
325 → Maximal A
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
326 Zorn-lemma {A} 0<A PO supP = zorn04 where
493
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 492
diff changeset
327 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
328 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
c03d80290855 total of B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 495
diff changeset
329 z01 {a} {b} A∋a A∋b (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO {me A∋b} {me A∋b} refl (IsStrictPartialOrder.trans PO {me A∋b} {me A∋a} {me A∋b} b<a a<b)
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
330 z02 : {x : Ordinal } → (ax : A ∋ * x ) → InFiniteIChain A ax → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
331 z02 {x} ax ic = zc5 ic where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
332 FC : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
333 FC = IChainSet {A} (me ax)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
334 zc6 : InFiniteIChain A ax → ¬ SUP A FC
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
335 zc6 inf = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
336 FC-is-total : IsTotalOrderSet FC
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
337 FC-is-total = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
338 FC⊆A : FC ⊆ A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
339 FC⊆A = record { incl = λ {x} lt → proj1 lt }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
340 zc5 : InFiniteIChain A ax → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
341 zc5 x = zc6 x ( supP FC FC⊆A FC-is-total )
478
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 477
diff changeset
342 -- ZChain is not compatible with the SUP condition
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
343 ind : (x : Ordinal) → ((y : Ordinal) → y o< x → ZChain A y ∨ Maximal A )
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
344 → ZChain A x ∨ Maximal A
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
345 ind x prev with Oprev-p x
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
346 ... | yes op with ODC.∋-p O A (* x)
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
347 ... | no ¬Ax = zc1 where
476
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 475
diff changeset
348 -- we have previous ordinal and ¬ A ∋ x, use previous Zchain
471
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 470
diff changeset
349 px = Oprev.oprev op
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
350 zc1 : ZChain A x ∨ Maximal A
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
351 zc1 with prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc)
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
352 ... | case2 x = case2 x -- we have the Maximal
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
353 ... | case1 z with trio< x (& (ZChain.max z))
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
354 ... | tri< a ¬b ¬c = case1 record { max = ZChain.max z ; y<max = a }
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
355 ... | tri≈ ¬a b ¬c = {!!} -- x = max so ¬ A ∋ max
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
356 ... | tri> ¬a ¬b c = {!!} -- can't happen
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
357 ... | yes ax = zc4 where -- we have previous ordinal and A ∋ x
472
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 471
diff changeset
358 px = Oprev.oprev op
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
359 zc1 : OSup> A (subst (OD.def (od A)) (sym &iso) ax) → ZChain A x ∨ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
360 zc1 os with prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc)
498
8ec0b88b022f zorn-case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 497
diff changeset
361 ... | case2 x = case2 x
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
362 ... | case1 x = {!!}
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
363 zc4 : ZChain A x ∨ Maximal A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
364 zc4 with Zorn-lemma-3case 0<A PO (me ax)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
365 ... | case1 y>x = zc1 y>x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
366 ... | case2 (case1 x) = case2 x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
367 ... | case2 (case2 x) = ⊥-elim (zc5 x) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
368 FC : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
369 FC = IChainSet {A} (me ax)
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
370 B : InFiniteIChain A ax → HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
371 B ifc = InFCSet A ax ifc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
372 zc6 : (ifc : InFiniteIChain A ax ) → ¬ SUP A (B ifc)
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
373 zc6 = {!!}
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
374 FC-is-total : (ifc : InFiniteIChain A ax) → IsTotalOrderSet (B ifc)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
375 FC-is-total ifc = ChainClosure-is-total A ax PO ifc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
376 B⊆A : (ifc : InFiniteIChain A ax) → B ifc ⊆ A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
377 B⊆A = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
378 ifc : InFiniteIChain A (subst (OD.def (od A)) (sym &iso) ax) → InFiniteIChain A ax
512
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
379 ifc record { chain<x = chain<x ; c-infinite = c-infinite } = record { chain<x = {!!} ; c-infinite = {!!} } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
380 ifc01 : {!!} -- me (subst (OD.def (od A)) (sym &iso) ax)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
381 ifc01 = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
382 -- (y : Ordinal) → odef (IChainSet (me (subst (OD.def (od A)) (sym &iso) ax))) y → y o< & (* x₁)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
383 -- (y : Ordinal) → odef (IChainSet (me ax)) y → y o< x₁
503
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 502
diff changeset
384 zc5 : InFiniteIChain A (subst (OD.def (od A)) (sym &iso) ax) → ⊥
511
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 510
diff changeset
385 zc5 x = zc6 (ifc x) ( supP (B (ifc x)) (B⊆A (ifc x)) (FC-is-total (ifc x) ))
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
386 ind x prev | no ¬ox with trio< (& A) x --- limit ordinal case
483
ed29002a02b6 zorn again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 482
diff changeset
387 ... | tri< a ¬b ¬c = {!!} where
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
388 zc1 : ZChain A (& A)
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
389 zc1 with prev (& A) a
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
390 ... | t = {!!}
483
ed29002a02b6 zorn again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 482
diff changeset
391 ... | tri≈ ¬a b ¬c = {!!} where
478
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 477
diff changeset
392 ... | tri> ¬a ¬b c with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 477
diff changeset
393 ... | no ¬Ax = {!!} where
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
394 ... | yes ax = {!!}
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
395 zorn03 : (x : Ordinal) → ZChain A x ∨ Maximal A
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
396 zorn03 x = TransFinite ind x
515
5faeae7cfe22 ε-induction does not work on Zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 514
diff changeset
397 zorn07 : (x : HOD) → (ax : A ∋ x ) → OSup> A (d→∋ A ax) ∨ Maximal A ∨ InFiniteIChain A (d→∋ A ax)
5faeae7cfe22 ε-induction does not work on Zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 514
diff changeset
398 zorn07 x = ε-induction (λ {x} prev ax → Zorn-lemma-3case 0<A PO (me {A} {x} ax ) ) x
5faeae7cfe22 ε-induction does not work on Zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 514
diff changeset
399 zorn05 : Maximal A
5faeae7cfe22 ε-induction does not work on Zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 514
diff changeset
400 zorn05 with zorn07 A {!!}
5faeae7cfe22 ε-induction does not work on Zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 514
diff changeset
401 ... | case1 chain = {!!}
5faeae7cfe22 ε-induction does not work on Zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 514
diff changeset
402 ... | case2 (case1 m) = m
5faeae7cfe22 ε-induction does not work on Zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 514
diff changeset
403 ... | case2 (case2 chain) = {!!}
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
404 zorn04 : Maximal A
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
405 zorn04 with zorn03 (& A)
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
406 ... | case1 chain = ⊥-elim ( o<> (c<→o< {ZChain.max chain} {A} (ZChain.A∋max chain)) (ZChain.y<max chain) )
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
407 ... | case2 m = m
464
5acf6483a9e3 Zorn lemma start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 447
diff changeset
408
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
409 -- _⊆'_ : ( A B : HOD ) → Set n
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
410 -- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x
482
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 481
diff changeset
411
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
412 -- MaximumSubset : {L P : HOD}
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
413 -- → o∅ o< & L → o∅ o< & P → P ⊆ L
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
414 -- → IsPartialOrderSet P _⊆'_
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
415 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ )
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
416 -- → Maximal P (_⊆'_)
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
417 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆'_} 0<P PO SP