Mercurial > hg > Members > kono > Proof > ZF-in-agda
changeset 43:0d9b9db14361
equalitu and internal parametorisity
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Fri, 24 May 2019 22:22:16 +0900 |
parents | 4d5fc6381546 |
children | fcac01485f32 |
files | ordinal-definable.agda ordinal.agda |
diffstat | 2 files changed, 66 insertions(+), 35 deletions(-) [+] |
line wrap: on
line diff
--- a/ordinal-definable.agda Fri May 24 10:28:02 2019 +0900 +++ b/ordinal-definable.agda Fri May 24 22:22:16 2019 +0900 @@ -34,6 +34,28 @@ _c<_ : { n : Level } → ( a x : OD {n} ) → Set n x c< a = a ∋ x +-- _=='_ : {n : Level} → Set (suc n) -- Rel (OD {n}) (suc n) +-- _=='_ {n} = ( a b : OD {n} ) → ( ∀ { x : OD {n} } → a ∋ x → b ∋ x ) ∧ ( ∀ { x : OD {n} } → a ∋ x → b ∋ x ) + +record _==_ {n : Level} ( a b : OD {n} ) : Set n where + field + eq→ : ∀ { x : Ordinal {n} } → def a x → def b x + eq← : ∀ { x : Ordinal {n} } → def b x → def a x + +id : {n : Level} {A : Set n} → A → A +id x = x + +eq-refl : {n : Level} { x : OD {n} } → x == x +eq-refl {n} {x} = record { eq→ = id ; eq← = id } + +open _==_ + +eq-sym : {n : Level} { x y : OD {n} } → x == y → y == x +eq-sym eq = record { eq→ = eq← eq ; eq← = eq→ eq } + +eq-trans : {n : Level} { x y z : OD {n} } → x == y → y == z → x == z +eq-trans x=y y=z = record { eq→ = λ t → eq→ y=z ( eq→ x=y t) ; eq← = λ t → eq← x=y ( eq← y=z t) } + _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n) a c≤ b = (a ≡ b) ∨ ( b ∋ a ) @@ -136,8 +158,8 @@ open Ordinal -∋-subst : {n : Level} {X Y x y : OD {suc n} } → X ≡ x → Y ≡ y → X ∋ Y → x ∋ y -∋-subst refl refl x = x +-- ∋-subst : {n : Level} {X Y x y : OD {suc n} } → X ≡ x → Y ≡ y → X ∋ Y → x ∋ y +-- ∋-subst refl refl x = x -- ∅77 : {n : Level} → (x : OD {suc n} ) → ¬ ( ord→od (o∅ {suc n}) ∋ x ) -- ∅77 {n} x lt = {!!} where @@ -145,19 +167,26 @@ ∅7' : {n : Level} → ord→od (o∅ {n}) ≡ od∅ {n} ∅7' {n} = cong ( λ k → record { def = k }) ( ∅-base-def ) where -∅7 : {n : Level} → ( x : OD {n} ) → od→ord x ≡ o∅ {n} → x ≡ od∅ {n} -∅7 {n} x eq = trans ( trans (sym oiso)( cong ( λ k → ord→od k ) eq ) ) ∅7' +∅7 : {n : Level} → ( x : OD {n} ) → od→ord x ≡ o∅ {n} → x == od∅ {n} +∅7 {n} x eq = record { eq→ = e1 ; eq← = e2 } where + e0 : {y : Ordinal {n}} → y o< o∅ {n} → def od∅ y + e0 {y} (case1 ()) + e0 {y} (case2 ()) + e1 : {y : Ordinal {n}} → def x y → def od∅ y + e1 {y} y<x = e0 ( o<-subst ( c<→o< {n} {x} y<x ) refl {!!} ) + e2 : {y : Ordinal {n}} → def od∅ y → def x y + e2 {y} (lift ()) -∅9 : {n : Level} → (x : OD {n} ) → ¬ x ≡ od∅ → o∅ o< od→ord x +∅9 : {n : Level} → (x : OD {n} ) → ¬ x == od∅ → o∅ o< od→ord x ∅9 x not = ∅5 ( od→ord x) lemma where lemma : ¬ od→ord x ≡ o∅ lemma eq = not ( ∅7 x eq ) -OD→ZF : {n : Level} → ZF {suc n} {suc n} +OD→ZF : {n : Level} → ZF {suc n} {n} OD→ZF {n} = record { ZFSet = OD {n} - ; _∋_ = λ a x → Lift (suc n) ( a ∋ x ) - ; _≈_ = _≡_ + ; _∋_ = _∋_ + ; _≈_ = _==_ ; ∅ = od∅ ; _,_ = _,_ ; Union = Union @@ -169,11 +198,10 @@ } where Replace : OD {n} → (OD {n} → OD {n} ) → OD {n} Replace X ψ = sup-od ψ - Select : OD {n} → (OD {n} → Set (suc n) ) → OD {n} + Select : OD {n} → (OD {n} → Set n ) → OD {n} Select X ψ = record { def = λ x → select ( ord→od x ) } where select : OD {n} → Set n - select x with ψ x - ... | t = Lift n ⊤ + select x = ψ x _,_ : OD {n} → OD {n} → OD {n} x , y = record { def = λ z → ( (z ≡ od→ord x ) ∨ ( z ≡ od→ord y )) } Union : OD {n} → OD {n} @@ -186,15 +214,15 @@ _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set n _⊆_ A B {x} = A ∋ x → B ∋ x _∩_ : ( A B : ZFSet ) → ZFSet - A ∩ B = Select (A , B) ( λ x → (Lift (suc n) ( A ∋ x )) ∧ (Lift (suc n) ( B ∋ x ) )) + A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ (B ∋ x) ) _∪_ : ( A B : ZFSet ) → ZFSet - A ∪ B = Select (A , B) ( λ x → (Lift (suc n) ( A ∋ x )) ∨ (Lift (suc n) ( B ∋ x ) )) + A ∪ B = Select (A , B) ( λ x → (A ∋ x) ∨ ( B ∋ x ) ) infixr 200 _∈_ infixr 230 _∩_ _∪_ infixr 220 _⊆_ - isZF : IsZF (OD {n}) (λ a x → Lift (suc n) ( a ∋ x )) _≡_ od∅ _,_ Union Power Select Replace (record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } }) + isZF : IsZF (OD {n}) _∋_ _==_ od∅ _,_ Union Power Select Replace (record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } }) isZF = record { - isEquivalence = record { refl = refl ; sym = sym ; trans = trans } + isEquivalence = record { refl = eq-refl ; sym = eq-sym; trans = eq-trans } ; pair = pair ; union→ = {!!} ; union← = {!!} @@ -211,31 +239,31 @@ } where open _∧_ open Minimumo - pair : (A B : OD {n} ) → Lift (suc n) ((A , B) ∋ A) ∧ Lift (suc n) ((A , B) ∋ B) - proj1 (pair A B ) = lift ( case1 refl ) - proj2 (pair A B ) = lift ( case2 refl ) - empty : (x : OD {n} ) → ¬ Lift (suc n) (od∅ ∋ x) - empty x (lift (lift ())) - union→ : (X x y : OD {n} ) → Lift (suc n) (X ∋ x) → Lift (suc n) (x ∋ y) → Lift (suc n) (Union X ∋ y) - union→ X x y (lift X∋x) (lift x∋y) = lift {!!} where + pair : (A B : OD {n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B) + proj1 (pair A B ) = case1 refl + proj2 (pair A B ) = case2 refl + empty : (x : OD {n} ) → ¬ (od∅ ∋ x) + empty x () + union→ : (X x y : OD {n} ) → (X ∋ x) → (x ∋ y) → (Union X ∋ y) + union→ X x y X∋x x∋y = {!!} where lemma : {z : Ordinal {n} } → def X z → z ≡ od→ord y lemma {z} X∋z = {!!} - minord : (x : OD {n} ) → ¬ x ≡ od∅ → Minimumo (od→ord x) + minord : (x : OD {n} ) → ¬ (x == od∅ )→ Minimumo (od→ord x) minord x not = ominimal (od→ord x) (∅9 x not) - minimul : (x : OD {n} ) → ¬ x ≡ od∅ → OD {n} + minimul : (x : OD {n} ) → ¬ (x == od∅ )→ OD {n} minimul x not = ord→od ( mino (minord x not)) - minimul<x : (x : OD {n} ) → (not : ¬ x ≡ od∅ ) → x ∋ minimul x not + minimul<x : (x : OD {n} ) → (not : ¬ x == od∅ ) → x ∋ minimul x not minimul<x x not = lemma0 (min<x (minord x not)) where lemma0 : mino (minord x not) o< (od→ord x) → def x (od→ord (ord→od (mino (minord x not)))) lemma0 m<x = def-subst {n} {ord→od (od→ord x)} {od→ord (ord→od (mino (minord x not)))} (o<→c< m<x) oiso refl - regularity : (x : OD) → (not : ¬ x ≡ od∅ ) → - Lift (suc n) (x ∋ minimul x not ) ∧ - (Select x (λ x₁ → Lift (suc n) ( minimul x not ∋ x₁) ∧ Lift (suc n) (x ∋ x₁)) ≡ od∅) - proj1 ( regularity x non ) = lift ( minimul<x x non ) - proj2 ( regularity x non ) = cong ( λ k → record { def = k } ) ( extensionality ( λ y → lemma0 y) ) where - lemma : ( z : Ordinal {n} ) → def (Select x (λ y → Lift (suc n) (minimul x non ∋ y) ∧ Lift (suc n) (x ∋ y))) z - lemma z with (minimul x non ∋ (ord→od z)) | x ∋ (ord→od z) - ... | s | t = {!!} - lemma0 : ( z : Ordinal {n} ) → def (Select x (λ y → Lift (suc n) (minimul x non ∋ y) ∧ Lift (suc n) (x ∋ y))) z ≡ Lift n ⊥ - lemma0 = {!!} + regularity : (x : OD) (not : ¬ (x == od∅)) → (x ∋ minimul x not) ∧ + (Select (minimul x not , x) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅) + -- regularity : (x : OD) → (not : ¬ x == od∅ ) → + -- ((x ∋ minimul x not ) ∧ {!!} ) -- (Select x (λ x₁ → (( minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅))) + proj1 ( regularity x non ) = minimul<x x non + proj2 ( regularity x non ) = {!!} where -- cong ( λ k → record { def = k } ) ( extensionality ( λ y → lemma0 y) ) where + not-min : ( z : Ordinal {n} ) → ¬ ( def (Select x (λ y → (minimul x non ∋ y) ∧ (x ∋ y))) z ) + not-min z = {!!} + lemma0 : ( z : Ordinal {n} ) → def (Select x (λ y → (minimul x non ∋ y) ∧ (x ∋ y))) z ≡ Lift n ⊥ + lemma0 z = {!!}
--- a/ordinal.agda Fri May 24 10:28:02 2019 +0900 +++ b/ordinal.agda Fri May 24 22:22:16 2019 +0900 @@ -35,6 +35,9 @@ _o<_ : {n : Level} ( x y : Ordinal ) → Set n _o<_ x y = (lv x < lv y ) ∨ ( ord x d< ord y ) +o<-subst : {n : Level } {Z X z x : Ordinal {n}} → Z o< X → Z ≡ z → X ≡ x → z o< x +o<-subst df refl refl = df + open import Data.Nat.Properties open import Data.Empty open import Data.Unit using ( ⊤ )