Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate src/ODC.agda @ 1353:ddbc0726f9bb
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 18 Jun 2023 18:16:03 +0900 |
parents | 47d3cc596d68 |
children | 34f72406fcfd |
rev | line source |
---|---|
431 | 1 {-# OPTIONS --allow-unsolved-metas #-} |
2 open import Level | |
3 open import Ordinals | |
4 module ODC {n : Level } (O : Ordinals {n} ) where | |
5 | |
1120 | 6 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) |
431 | 7 open import Relation.Binary.PropositionalEquality |
1120 | 8 open import Data.Nat.Properties |
431 | 9 open import Data.Empty |
10 open import Relation.Nullary | |
11 open import Relation.Binary | |
12 open import Relation.Binary.Core | |
13 | |
14 import OrdUtil | |
15 open import logic | |
16 open import nat | |
17 import OD | |
18 import ODUtil | |
19 | |
20 open inOrdinal O | |
21 open OD O | |
22 open OD.OD | |
23 open OD._==_ | |
24 open ODAxiom odAxiom | |
25 open ODUtil O | |
26 | |
27 open Ordinals.Ordinals O | |
28 open Ordinals.IsOrdinals isOrdinal | |
1300 | 29 -- open Ordinals.IsNext isNext |
431 | 30 open OrdUtil O |
31 | |
32 | |
33 open HOD | |
34 | |
35 open _∧_ | |
36 | |
1120 | 37 postulate |
431 | 38 -- mimimul and x∋minimal is an Axiom of choice |
1120 | 39 minimal : (x : HOD ) → ¬ (x =h= od∅ )→ HOD |
431 | 40 -- this should be ¬ (x =h= od∅ )→ ∃ ox → x ∋ Ord ox ( minimum of x ) |
41 x∋minimal : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → odef x ( & ( minimal x ne ) ) | |
467 | 42 -- minimality (proved by ε-induction with LEM) |
431 | 43 minimal-1 : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → (y : HOD ) → ¬ ( odef (minimal x ne) (& y)) ∧ (odef x (& y) ) |
44 | |
45 | |
46 -- | |
47 -- Axiom of choice in intutionistic logic implies the exclude middle | |
48 -- https://plato.stanford.edu/entries/axiom-choice/#AxiChoLog | |
49 -- | |
50 | |
51 pred-od : ( p : Set n ) → HOD | |
52 pred-od p = record { od = record { def = λ x → (x ≡ o∅) ∧ p } ; | |
1120 | 53 odmax = osuc o∅; <odmax = λ x → subst (λ k → k o< osuc o∅) (sym (proj1 x)) <-osuc } |
431 | 54 |
55 ppp : { p : Set n } { a : HOD } → pred-od p ∋ a → p | |
56 ppp {p} {a} d = proj2 d | |
57 | |
58 p∨¬p : ( p : Set n ) → p ∨ ( ¬ p ) -- assuming axiom of choice | |
59 p∨¬p p with is-o∅ ( & (pred-od p )) | |
60 p∨¬p p | yes eq = case2 (¬p eq) where | |
1120 | 61 ps = pred-od p |
62 eqo∅ : ps =h= od∅ → & ps ≡ o∅ | |
63 eqo∅ eq = subst (λ k → & ps ≡ k) ord-od∅ ( cong (λ k → & k ) (==→o≡ eq)) | |
431 | 64 lemma : ps =h= od∅ → p → ⊥ |
65 lemma eq p0 = ¬x<0 {& ps} (eq→ eq record { proj1 = eqo∅ eq ; proj2 = p0 } ) | |
66 ¬p : (& ps ≡ o∅) → p → ⊥ | |
67 ¬p eq = lemma ( subst₂ (λ j k → j =h= k ) *iso o∅≡od∅ ( o≡→== eq )) | |
68 p∨¬p p | no ¬p = case1 (ppp {p} {minimal ps (λ eq → ¬p (eqo∅ eq))} lemma) where | |
1120 | 69 ps = pred-od p |
70 eqo∅ : ps =h= od∅ → & ps ≡ o∅ | |
71 eqo∅ eq = subst (λ k → & ps ≡ k) ord-od∅ ( cong (λ k → & k ) (==→o≡ eq)) | |
72 lemma : ps ∋ minimal ps (λ eq → ¬p (eqo∅ eq)) | |
431 | 73 lemma = x∋minimal ps (λ eq → ¬p (eqo∅ eq)) |
74 | |
1120 | 75 decp : ( p : Set n ) → Dec p -- assuming axiom of choice |
431 | 76 decp p with p∨¬p p |
77 decp p | case1 x = yes x | |
78 decp p | case2 x = no x | |
79 | |
1120 | 80 ∋-p : (A x : HOD ) → Dec ( A ∋ x ) |
431 | 81 ∋-p A x with p∨¬p ( A ∋ x ) -- LEM |
82 ∋-p A x | case1 t = yes t | |
83 ∋-p A x | case2 t = no (λ x → t x) | |
84 | |
1120 | 85 double-neg-elim : {A : Set n} → ¬ ¬ A → A -- we don't have this in intutionistic logic |
86 double-neg-elim {A} notnot with decp A -- assuming axiom of choice | |
431 | 87 ... | yes p = p |
88 ... | no ¬p = ⊥-elim ( notnot ¬p ) | |
89 | |
860
105f8d6c51fb
no-extension on immidate ordinal passed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
477
diff
changeset
|
90 or-exclude : {A B : Set n} → A ∨ B → A ∨ ( (¬ A) ∧ B ) |
105f8d6c51fb
no-extension on immidate ordinal passed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
477
diff
changeset
|
91 or-exclude {A} {B} ab with p∨¬p A |
105f8d6c51fb
no-extension on immidate ordinal passed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
477
diff
changeset
|
92 or-exclude {A} {B} (case1 a) | case1 a0 = case1 a |
105f8d6c51fb
no-extension on immidate ordinal passed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
477
diff
changeset
|
93 or-exclude {A} {B} (case1 a) | case2 ¬a = ⊥-elim ( ¬a a ) |
105f8d6c51fb
no-extension on immidate ordinal passed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
477
diff
changeset
|
94 or-exclude {A} {B} (case2 b) | case1 a = case1 a |
105f8d6c51fb
no-extension on immidate ordinal passed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
477
diff
changeset
|
95 or-exclude {A} {B} (case2 b) | case2 ¬a = case2 ⟪ ¬a , b ⟫ |
105f8d6c51fb
no-extension on immidate ordinal passed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
477
diff
changeset
|
96 |
1120 | 97 -- record By-contradiction (A : Set n) (B : A → Set n) : Set (suc n) where |
98 -- field | |
99 -- a : A | |
100 -- b : B a | |
101 -- | |
102 -- by-contradiction : {A : Set n} {B : A → Set n} → ¬ ( (a : A ) → ¬ B a ) → By-contradiction A B | |
103 -- by-contradiction {A} {B} not with p∨¬p A | |
104 -- ... | case2 ¬a = ⊥-elim (not (λ a → ⊥-elim (¬a a ))) | |
105 -- ... | case1 a with p∨¬p (B a) | |
106 -- ... | case1 b = record { a = a ; b = b } | |
107 -- ... | case2 ¬b = ⊥-elim ( not ( λ a b → ⊥-elim ( ¬b ? ))) | |
108 -- | |
431 | 109 power→⊆ : ( A t : HOD) → Power A ∋ t → t ⊆ A |
1096 | 110 power→⊆ A t PA∋t tx = subst (λ k → odef A k ) &iso ( power→ A t PA∋t (subst (λ k → odef t k) (sym &iso) tx ) ) |
431 | 111 |
447 | 112 power-∩ : { A x y : HOD } → Power A ∋ x → Power A ∋ y → Power A ∋ ( x ∩ y ) |
113 power-∩ {A} {x} {y} ax ay = power← A (x ∩ y) p01 where | |
114 p01 : {z : HOD} → (x ∩ y) ∋ z → A ∋ z | |
1096 | 115 p01 {z} xyz = power→ A x ax (proj1 xyz ) |
447 | 116 |
431 | 117 OrdP : ( x : Ordinal ) ( y : HOD ) → Dec ( Ord x ∋ y ) |
118 OrdP x y with trio< x (& y) | |
119 OrdP x y | tri< a ¬b ¬c = no ¬c | |
120 OrdP x y | tri≈ ¬a refl ¬c = no ( o<¬≡ refl ) | |
121 OrdP x y | tri> ¬a ¬b c = yes c | |
122 | |
123 open import zfc | |
124 | |
125 HOD→ZFC : ZFC | |
1120 | 126 HOD→ZFC = record { |
127 ZFSet = HOD | |
128 ; _∋_ = _∋_ | |
129 ; _≈_ = _=h=_ | |
431 | 130 ; ∅ = od∅ |
131 ; Select = Select | |
132 ; isZFC = isZFC | |
133 } where | |
134 -- infixr 200 _∈_ | |
135 -- infixr 230 _∩_ _∪_ | |
136 isZFC : IsZFC (HOD ) _∋_ _=h=_ od∅ Select | |
137 isZFC = record { | |
138 choice-func = choice-func ; | |
139 choice = choice | |
140 } where | |
141 -- Axiom of choice ( is equivalent to the existence of minimal in our case ) | |
1120 | 142 -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] |
431 | 143 choice-func : (X : HOD ) → {x : HOD } → ¬ ( x =h= od∅ ) → ( X ∋ x ) → HOD |
144 choice-func X {x} not X∋x = minimal x not | |
1120 | 145 choice : (X : HOD ) → {A : HOD } → ( X∋A : X ∋ A ) → (not : ¬ ( A =h= od∅ )) → A ∋ choice-func X not X∋A |
431 | 146 choice X {A} X∋A not = x∋minimal A not |
147 |