AC.agda 6.1 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193
  1. {-# OPTIONS --no-termination-check #-}
  2. module AC where
  3. import Nat
  4. import Bool
  5. import List
  6. import Fin
  7. import Logic
  8. import Vec
  9. import EqProof
  10. open Nat hiding (_<_) renaming (_==_ to _=Nat=_)
  11. open Bool
  12. open List hiding (module Eq)
  13. open Fin renaming (_==_ to _=Fin=_)
  14. open Logic
  15. open Vec
  16. infix 20 _○_
  17. infix 10 _≡_
  18. ForAll : {A : Set}(n : Nat) -> (Vec n A -> Set) -> Set
  19. ForAll zero F = F ε
  20. ForAll {A} (suc n) F = (x : A) -> ForAll _ \xs -> F (x ∷ xs)
  21. apply : {n : Nat}{A : Set}(F : Vec n A -> Set) -> ForAll n F -> (xs : Vec n A) -> F xs
  22. apply {zero} F t (vec vnil) = t
  23. apply {suc n} F f (vec (vcons x xs)) = apply _ (f x) xs
  24. lambda : {n : Nat}{A : Set}(F : Vec n A -> Set) -> ((xs : Vec n A) -> F xs) -> ForAll n F
  25. lambda {zero } F f = f ε
  26. lambda {suc n} F f = \x -> lambda _ (\xs -> f (x ∷ xs))
  27. data Expr (n : Nat) : Set where
  28. zro : Expr n
  29. var : Fin n -> Expr n
  30. _○_ : Expr n -> Expr n -> Expr n
  31. data Theorem (n : Nat) : Set where
  32. _≡_ : Expr n -> Expr n -> Theorem n
  33. theorem : (n : Nat) -> ({m : Nat} -> ForAll {Expr m} n \_ -> Theorem m) -> Theorem n
  34. theorem n thm = apply _ thm (map var (fzeroToN-1 n))
  35. module Provable where
  36. NF : Nat -> Set
  37. NF n = List (Fin n)
  38. infix 12 _⊕_
  39. _⊕_ : {n : Nat} -> NF n -> NF n -> NF n
  40. [] ⊕ ys = ys
  41. x :: xs ⊕ [] = x :: xs
  42. x :: xs ⊕ y :: ys = if x < y
  43. then x :: (xs ⊕ y :: ys)
  44. else y :: (x :: xs ⊕ ys)
  45. normalise : {n : Nat} -> Expr n -> NF n
  46. normalise zro = []
  47. normalise (var n) = n :: []
  48. normalise (a ○ b) = normalise a ⊕ normalise b
  49. infix 30 _↓
  50. _↓ : {n : Nat} -> NF n -> Expr n
  51. (i :: is) ↓ = var i ○ is ↓
  52. [] ↓ = zro
  53. infix 10 _=Expr=_ _=NF=_
  54. _=NF=_ : {n : Nat} -> NF n -> NF n -> Bool
  55. _=NF=_ = ListEq._==_
  56. where
  57. module ListEq = List.Eq _=Fin=_
  58. substNF : {n : Nat}{xs ys : NF n}(P : NF n -> Set) -> IsTrue (xs =NF= ys) -> P xs -> P ys
  59. substNF = List.Subst.subst _=Fin=_ (subst {_})
  60. _=Expr=_ : {n : Nat} -> Expr n -> Expr n -> Bool
  61. a =Expr= b = normalise a =NF= normalise b
  62. provable : {n : Nat} -> Theorem n -> Bool
  63. provable (a ≡ b) = a =Expr= b
  64. module Semantics
  65. {A : Set}
  66. (_==_ : A -> A -> Set)
  67. (_*_ : A -> A -> A)
  68. (one : A)
  69. (refl : {x : A} -> x == x)
  70. (sym : {x y : A} -> x == y -> y == x)
  71. (trans : {x y z : A} -> x == y -> y == z -> x == z)
  72. (idL : {x : A} -> (one * x) == x)
  73. (idR : {x : A} -> (x * one) == x)
  74. (comm : {x y : A} -> (x * y) == (y * x))
  75. (assoc : {x y z : A} -> (x * (y * z)) == ((x * y) * z))
  76. (congL : {x y z : A} -> y == z -> (x * y) == (x * z))
  77. (congR : {x y z : A} -> x == y -> (x * z) == (y * z))
  78. where
  79. open Provable
  80. module EqP = EqProof _==_ refl trans
  81. open EqP
  82. expr[_] : {n : Nat} -> Expr n -> Vec n A -> A
  83. expr[ zro ] ρ = one
  84. expr[ var i ] ρ = ρ ! i
  85. expr[ a ○ b ] ρ = expr[ a ] ρ * expr[ b ] ρ
  86. eq[_] : {n : Nat} -> Theorem n -> Vec n A -> Set
  87. eq[ a ≡ b ] ρ = expr[ a ] ρ == expr[ b ] ρ
  88. data CantProve (A : Set) : Set where
  89. no-proof : CantProve A
  90. Prf : {n : Nat} -> Theorem n -> Bool -> Set
  91. Prf thm true = ForAll _ \ρ -> eq[ thm ] ρ
  92. Prf thm false = CantProve (Prf thm true)
  93. Proof : {n : Nat} -> Theorem n -> Set
  94. Proof thm = Prf thm (provable thm)
  95. lem0 : {n : Nat} -> (xs ys : NF n) -> (ρ : Vec n A) ->
  96. eq[ xs ↓ ○ ys ↓ ≡ (xs ⊕ ys) ↓ ] ρ
  97. lem0 [] ys ρ = idL
  98. lem0 (x :: xs) [] ρ = idR
  99. lem0 (x :: xs) (y :: ys) ρ = if' x < y then less else more
  100. where
  101. lhs = (var x ○ xs ↓) ○ (var y ○ ys ↓)
  102. lbranch = x :: (xs ⊕ y :: ys)
  103. rbranch = y :: (x :: xs ⊕ ys)
  104. P = \z -> eq[ lhs ≡ (if z then lbranch else rbranch) ↓ ] ρ
  105. less : IsTrue (x < y) -> _
  106. less x<y = BoolEq.subst {true}{x < y} P x<y
  107. (spine (lem0 xs (y :: ys) ρ))
  108. where
  109. spine : forall {x' xs' y' ys' zs} h -> _
  110. spine {x'}{xs'}{y'}{ys'}{zs} h =
  111. eqProof> ((x' * xs') * (y' * ys'))
  112. === (x' * (xs' * (y' * ys'))) by sym assoc
  113. === (x' * zs) by congL h
  114. more : IsFalse (x < y) -> _
  115. more x>=y = BoolEq.subst {false}{x < y} P x>=y
  116. (spine (lem0 (x :: xs) ys ρ))
  117. where
  118. spine : forall {x' xs' y' ys' zs} h -> _
  119. spine {x'}{xs'}{y'}{ys'}{zs} h =
  120. eqProof> ((x' * xs') * (y' * ys'))
  121. === ((y' * ys') * (x' * xs')) by comm
  122. === (y' * (ys' * (x' * xs'))) by sym assoc
  123. === (y' * ((x' * xs') * ys')) by congL comm
  124. === (y' * zs) by congL h
  125. lem1 : {n : Nat} -> (e : Expr n) -> (ρ : Vec n A) -> eq[ e ≡ normalise e ↓ ] ρ
  126. lem1 zro ρ = refl
  127. lem1 (var i) ρ = sym idR
  128. lem1 (a ○ b) ρ = trans step1 (trans step2 step3)
  129. where
  130. step1 : eq[ a ○ b ≡ normalise a ↓ ○ b ] ρ
  131. step1 = congR (lem1 a ρ)
  132. step2 : eq[ normalise a ↓ ○ b ≡ normalise a ↓ ○ normalise b ↓ ] ρ
  133. step2 = congL (lem1 b ρ)
  134. step3 : eq[ normalise a ↓ ○ normalise b ↓ ≡ (normalise a ⊕ normalise b) ↓ ] ρ
  135. step3 = lem0 (normalise a) (normalise b) ρ
  136. lem2 : {n : Nat} -> (xs ys : NF n) -> (ρ : Vec n A) -> IsTrue (xs =NF= ys) -> eq[ xs ↓ ≡ ys ↓ ] ρ
  137. lem2 xs ys ρ eq = substNF {_}{xs}{ys} (\z -> eq[ xs ↓ ≡ z ↓ ] ρ) eq refl
  138. prove : {n : Nat} -> (thm : Theorem n) -> Proof thm
  139. prove thm = proof (provable thm) thm (\h -> h)
  140. where
  141. proof : {n : Nat}(valid : Bool)(thm : Theorem n) -> (IsTrue valid -> IsTrue (provable thm)) -> Prf thm valid
  142. proof false _ _ = no-proof
  143. proof true (a ≡ b) isValid = lambda eq[ a ≡ b ] \ρ ->
  144. trans (step-a ρ) (trans (step-ab ρ) (step-b ρ))
  145. where
  146. step-a : forall ρ -> eq[ a ≡ normalise a ↓ ] ρ
  147. step-a ρ = lem1 a ρ
  148. step-b : forall ρ -> eq[ normalise b ↓ ≡ b ] ρ
  149. step-b ρ = sym (lem1 b ρ)
  150. step-ab : forall ρ -> eq[ normalise a ↓ ≡ normalise b ↓ ] ρ
  151. step-ab ρ = lem2 (normalise a) (normalise b) ρ (isValid tt)