Mercurial > hg > Members > kono > Proof > ZF-in-agda
diff HOD.agda @ 166:ea0e7927637a
use double negation
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Wed, 17 Jul 2019 10:52:31 +0900 |
parents | d16b8bf29f4f |
children | 4724f7be00e3 |
line wrap: on
line diff
--- a/HOD.agda Tue Jul 16 09:57:01 2019 +0900 +++ b/HOD.agda Wed Jul 17 10:52:31 2019 +0900 @@ -57,8 +57,8 @@ -- OD can be iso to a subset of Ordinal ( by means of Godel Set ) od→ord : {n : Level} → OD {n} → Ordinal {n} ord→od : {n : Level} → Ordinal {n} → OD {n} - c<→o< : {n : Level} {x y : OD {n} } → def y ( od→ord x ) → od→ord x o< od→ord y - oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x + c<→o< : {n : Level} {x y : OD {n} } → def y ( od→ord x ) → od→ord x o< od→ord y + oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x -- we should prove this in agda, but simply put here ==→o≡ : {n : Level} → { x y : OD {suc n} } → (x == y) → x ≡ y @@ -319,10 +319,9 @@ union→ X z u xx not = ⊥-elim ( not (od→ord u) ( record { proj1 = proj1 xx ; proj2 = subst ( λ k → def k (od→ord z)) (sym oiso) (proj2 xx) } )) union← : (X z : OD) (X∋z : Union X ∋ z) → ¬ ( (u : OD ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) - union← X z UX∋z = TransFiniteExists' _ lemma UX∋z where + union← X z UX∋z = TransFiniteExists _ lemma UX∋z where lemma : {y : Ordinal} → def X y ∧ def (ord→od y) (od→ord z) → ¬ ((u : OD) → ¬ (X ∋ u) ∧ (u ∋ z)) lemma {y} xx not = not (ord→od y) record { proj1 = subst ( λ k → def X k ) (sym diso) (proj1 xx ) ; proj2 = proj2 xx } - -- ( λ {y} xx → ψiso : {ψ : OD {suc n} → Set (suc n)} {x y : OD {suc n}} → ψ x → x ≡ y → ψ y ψiso {ψ} t refl = t @@ -377,22 +376,27 @@ lemma : od→ord (ZFSubset (Ord a) (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset (Ord a) (ord→od x))) lemma = sup-o< + -- double-neg-eilm : {n : Level } {A : Set n} → ¬ ¬ A → A + -- double-neg-eilm {n} {A} notnot = ⊥-elim (notnot (λ A → {!!} )) -- -- Every set in OD is a subset of Ordinals -- -- Power A = Replace (Def (Ord (od→ord A))) ( λ y → A ∩ y ) - power→ : ( A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → A ∋ x - power→ A t P∋t {x} t∋x = TransFiniteExists _ -- (λ y → (t == (A ∩ ord→od y))) - lemma4 lemma5 where + + -- we have oly double negation form because of the replacement axiom + -- + power→ : ( A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → ¬ ¬ (A ∋ x) + power→ A t P∋t {x} t∋x = TransFiniteExists _ lemma5 lemma4 where a = od→ord A lemma2 : ¬ ( (y : OD) → ¬ (t == (A ∩ y))) lemma2 = replacement→ (Def (Ord (od→ord A))) t P∋t - lemma3 : (y : OD) → t == ( A ∩ y ) → A ∋ x - lemma3 y eq = proj1 (eq→ eq t∋x) + lemma3 : (y : OD) → t == ( A ∩ y ) → ¬ ¬ (A ∋ x) + lemma3 y eq not = not (proj1 (eq→ eq t∋x)) lemma4 : ¬ ((y : Ordinal) → ¬ (t == (A ∩ ord→od y))) lemma4 not = lemma2 ( λ y not1 → not (od→ord y) (subst (λ k → t == ( A ∩ k )) (sym oiso) not1 )) - lemma5 : {y : Ordinal} → t == (A ∩ ord→od y) → def A (od→ord x) - lemma5 {y} eq = lemma3 (ord→od y) eq + lemma5 : {y : Ordinal} → t == (A ∩ ord→od y) → ¬ ¬ (def A (od→ord x)) + lemma5 {y} eq not = (lemma3 (ord→od y) eq) not + power← : (A t : OD) → ({x : OD} → (t ∋ x → A ∋ x)) → Power A ∋ t power← A t t→A = record { proj1 = lemma1 ; proj2 = lemma2 } where a = od→ord A