Mercurial > hg > Members > kono > Proof > automaton
view agda/halt.agda @ 176:728cd6f7bf56
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 14 Mar 2021 21:33:23 +0900 |
parents | bf50676c77af |
children | 375e1dcba6aa |
line wrap: on
line source
module halt where open import Level renaming ( zero to Zero ; suc to Suc ) open import Data.Nat open import Data.Maybe open import Data.List hiding ([_]) open import Data.Nat.Properties open import Relation.Nullary open import Data.Empty open import Data.Unit open import Relation.Binary.Core hiding (_⇔_) open import Relation.Binary.Definitions open import Relation.Binary.PropositionalEquality open import logic record HBijection {n m : Level} (R : Set n) (S : Set m) : Set (n Level.⊔ m) where field fun← : S → R fun→ : R → S fiso← : (x : R) → fun← ( fun→ x ) ≡ x -- normal bijection required below, but we don't need this to show the inconsistency -- fiso→ : (x : S ) → fun→ ( fun← x ) ≡ x injection : {n m : Level} (R : Set n) (S : Set m) (f : R → S ) → Set (n Level.⊔ m) injection R S f = (x y : R) → f x ≡ f y → x ≡ y open HBijection diag : {S : Set } (b : HBijection ( S → Bool ) S) → S → Bool diag b n = not (fun← b n n) diagonal : { S : Set } → ¬ HBijection ( S → Bool ) S diagonal {S} b = diagn1 (fun→ b (diag b) ) refl where diagn1 : (n : S ) → ¬ (fun→ b (diag b) ≡ n ) diagn1 n dn = ¬t=f (diag b n ) ( begin not (diag b n) ≡⟨⟩ not (not fun← b n n) ≡⟨ cong (λ k → not (k n) ) (sym (fiso← b _)) ⟩ not (fun← b (fun→ b (diag b)) n) ≡⟨ cong (λ k → not (fun← b k n) ) dn ⟩ not (fun← b n n) ≡⟨⟩ diag b n ∎ ) where open ≡-Reasoning postulate utm : List Bool → List Bool → Maybe Bool record TM : Set where field tm : List Bool → Maybe Bool encode : List Bool is-tm : utm encode ≡ tm open TM Halt1 : TM → List Bool → Bool -- ℕ → ( ℕ → Bool ) Halt1 = {!!} record Halt1-is-tm : Set where field htm : TM htm-is-Halt1 : (t : TM ) → (x : List Bool) → (Halt1 t x ≡ true) ⇔ ((tm htm (encode t ++ x)) ≡ just true) Halt2 : (tm : TM ) → List Bool -- ( ℕ → Bool ) → ℕ Halt2 tm = encode tm not-halt : {!!} not-halt = {!!} open _∧_ open ≡-Reasoning open import Axiom.Extensionality.Propositional postulate f-extensionality : { n : Level} → Axiom.Extensionality.Propositional.Extensionality n n open import Relation.Binary.HeterogeneousEquality as HE using (_≅_;refl ) renaming ( sym to ≅-sym ; trans to ≅-trans ; cong to ≅-cong ) record Halt : Set where field htm : TM is-halt : (t : TM ) → (x : List Bool ) → (tm htm (encode t ++ x) ≡ just true) ⇔ ((tm t x ≡ just true) ∨ (tm t x ≡ just false)) nhtm : TM nhtm-is-negation : (t : TM ) → (x : List Bool ) → (tm htm (encode t ++ x) ≡ just true) ⇔ (tm nhtm (encode t ++ x) ≡ nothing ) open Halt TNℕ : Halt → HBijection (TM → Bool) TM TNℕ = {!!} tm-cong : {x y : TM} → tm x ≡ tm y → encode x ≡ encode y → is-tm x ≅ is-tm y → x ≡ y tm-cong refl refl refl = refl tm-bij : HBijection TM (List Bool) tm-bij = record { fun← = λ x → record { tm = utm x ; encode = x ; is-tm = refl } ; fun→ = λ tm → encode tm ; fiso← = tmb1 } where tmb1 : (x : TM ) → record { tm = utm (encode x) ; encode = encode x ; is-tm = refl } ≡ x tmb1 x with is-tm x | inspect is-tm x ... | refl | record { eq = refl } = tm-cong (is-tm x) refl refl -- Halt1 (Halt2 x) ≡ x -- Halt2 (Halt2 y) ≡ y halting : ¬ Halt halting h with tm (htm h) (encode (nhtm h) ++ encode (nhtm h)) ... | just true = ¬t=f {!!} {!!} ... | nothing = {!!} ... | just false = {!!}