├── ClosedTautoSolver.agda ├── README.md └── TautoSolver.agda /ClosedTautoSolver.agda: -------------------------------------------------------------------------------- 1 | 2 | open import Prelude 3 | open import Container.List 4 | open import Tactic.Reflection 5 | open import Tactic.Deriving.Eq 6 | 7 | data P : Set where 8 | `⊤ : P 9 | `⊥ : P 10 | _`∧_ : P → P → P 11 | _`∨_ : P → P → P 12 | _`⇒_ : P → P → P 13 | 14 | unquoteDecl eqP = deriveEq eqP (quote P) 15 | 16 | Ctx = List P 17 | 18 | data _⊢_ (Γ : Ctx) : P → Set where 19 | var : ∀ {X} → X ∈ Γ → Γ ⊢ X 20 | lam : ∀ {X Y} → (X ∷ Γ) ⊢ Y → Γ ⊢ (X `⇒ Y) 21 | tt : Γ ⊢ `⊤ 22 | _,_ : ∀ {X Y} → Γ ⊢ X → Γ ⊢ Y → Γ ⊢ (X `∧ Y) 23 | inl : ∀ {X Y} → Γ ⊢ X → Γ ⊢ (X `∨ Y) 24 | inr : ∀ {X Y} → Γ ⊢ Y → Γ ⊢ (X `∨ Y) 25 | 26 | ⟦_⟧ : P → Set 27 | ⟦ `⊤ ⟧ = ⊤ 28 | ⟦ `⊥ ⟧ = ⊥ 29 | ⟦ X `∧ Y ⟧ = ⟦ X ⟧ × ⟦ Y ⟧ 30 | ⟦ X `∨ Y ⟧ = Either ⟦ X ⟧ ⟦ Y ⟧ 31 | ⟦ X `⇒ Y ⟧ = ⟦ X ⟧ → ⟦ Y ⟧ 32 | 33 | ⟦_⟧ctx : Ctx → Set 34 | ⟦_⟧ctx = All ⟦_⟧ 35 | 36 | ⊢-sound : ∀ {Γ P} → Γ ⊢ P → ⟦ Γ ⟧ctx → ⟦ P ⟧ 37 | ⊢-sound (var x) γ = lookup∈ γ x 38 | ⊢-sound tt γ = tt 39 | ⊢-sound (d₁ , d₂) γ = (⊢-sound d₁ γ) , (⊢-sound d₂ γ) 40 | ⊢-sound (inl d) γ = left (⊢-sound d γ) 41 | ⊢-sound (inr d) γ = right (⊢-sound d γ) 42 | ⊢-sound (lam d) γ z = ⊢-sound d (z ∷ γ) 43 | 44 | 45 | strengthn : Arg Type → Type → Term → TC Term 46 | strengthn dom a t = do 47 | t' ← newMeta a 48 | extendContext dom (unify t (weaken 1 t')) 49 | return t' 50 | 51 | {-# TERMINATING #-} 52 | parseP : Term → TC P 53 | parseP t = reduce t >>= λ where 54 | (def (quote ⊤) []) → return `⊤ 55 | (def (quote ⊥) []) → return `⊥ 56 | (def (quote Either) (_ ∷ _ ∷ (arg _ x) ∷ (arg _ y) ∷ [])) → do 57 | X ← parseP x 58 | Y ← parseP y 59 | return (X `∨ Y) 60 | (def (quote Σ) (_ ∷ _ ∷ ax@(arg _ x) ∷ (arg _ (lam _ (abs _ y))) ∷ [])) → do 61 | X ← parseP x 62 | y ← strengthn ax set! y <|> errorNotAProp t 63 | Y ← parseP y 64 | return (X `∧ Y) 65 | t@(pi ax@(arg _ x) (abs _ y)) → do 66 | X ← parseP x 67 | y ← strengthn ax set! y <|> errorNotAProp t 68 | Y ← parseP y 69 | return (X `⇒ Y) 70 | t → errorNotAProp t 71 | 72 | where 73 | errorNotAProp : ∀ {A} → Term → TC A 74 | errorNotAProp t = typeError (strErr "Parsing failed: " ∷ termErr t ∷ strErr "is not a proposition!" ∷ []) 75 | 76 | macro 77 | parse : Term → Term → TC ⊤ 78 | parse t goal = do 79 | T ← parseP t 80 | result ← quoteTC T 81 | unify goal result 82 | 83 | testType₁ : Set 84 | testType₁ = ⊤ × ⊥ → Either ⊥ ⊤ 85 | 86 | testParse₁ : parse (testType₁) ≡ ((`⊤ `∧ `⊥) `⇒ (`⊥ `∨ `⊤)) 87 | testParse₁ = refl 88 | 89 | assumption : (Γ : Ctx) (X : P) → Maybe (X ∈ Γ) 90 | assumption [] Y = nothing 91 | assumption (X ∷ Γ) Y with X == Y 92 | assumption (X ∷ Γ) .X | yes refl = just zero! 93 | assumption (X ∷ Γ) Y | no x = suc <$> assumption Γ Y 94 | 95 | solveTauto : (Γ : List P) (X : P) → Maybe (Γ ⊢ X) 96 | solveTauto Γ X = (var <$> assumption Γ X) <|> solveTautoAux X 97 | where 98 | solveTautoAux : (X : P) → Maybe (Γ ⊢ X) 99 | solveTautoAux = λ where 100 | `⊤ → return tt 101 | `⊥ → nothing 102 | (X `∧ Y) → do 103 | x ← solveTauto Γ X 104 | y ← solveTauto Γ Y 105 | return (x , y) 106 | (X `∨ Y) → (inl <$> solveTauto Γ X) <|> (inr <$> solveTauto Γ Y) 107 | (X `⇒ Y) → lam <$> solveTauto (X ∷ Γ) Y 108 | 109 | macro 110 | tauto : Term → TC ⊤ 111 | tauto goal = do 112 | X ← parseP =<< inferType goal 113 | case solveTauto [] X of λ where 114 | (just x) → let proof : ⟦ X ⟧ 115 | proof = ⊢-sound x [] 116 | in unify goal =<< quoteTC proof 117 | nothing → typeError (strErr "Failed to solve" ∷ []) 118 | 119 | test₁ : ⊤ 120 | test₁ = tauto 121 | 122 | test₂ : ⊥ → ⊥ 123 | test₂ = tauto 124 | 125 | test₃ : ⊥ → ⊤ → ⊥ × ⊤ 126 | test₃ = tauto 127 | 128 | test₄ : Either ⊥ ⊤ 129 | test₄ = tauto 130 | 131 | test₅ : ⊤ → Either ⊤ ⊥ 132 | test₅ = tauto 133 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | A sample tautology solver using Agda's reflection machinery -------------------------------------------------------------------------------- /TautoSolver.agda: -------------------------------------------------------------------------------- 1 | 2 | open import Prelude hiding (_>>=_) renaming (_>>=′_ to _>>=_) 3 | open import Container.List 4 | open import Container.Traversable 5 | open import Tactic.Reflection 6 | open import Tactic.Deriving.Eq 7 | 8 | module _ where 9 | 10 | module PreludeExtras where 11 | 12 | raise : ∀ {m} n → Fin m → Fin (n +N m) 13 | raise zero i = i 14 | raise (suc n) i = suc (raise n i) 15 | 16 | inject : ∀ {m} n → Fin m → Fin (m +N n) 17 | inject n zero = zero 18 | inject n (suc i) = suc (inject n i) 19 | 20 | module _ {a} {A : Set a} {{_ : Eq A}} where 21 | 22 | findIndex : ∀ {n} → Vec A n → A → Maybe (Fin n) 23 | findIndex [] a = nothing 24 | findIndex (x ∷ xs) a = case (x == a) of λ where 25 | (yes _) → just zero 26 | (no _ ) → suc <$> findIndex xs a 27 | 28 | open PreludeExtras 29 | 30 | 31 | data P (n : Nat) : Set where 32 | `⊤ : P n 33 | `⊥ : P n 34 | _`∧_ : P n → P n → P n 35 | _`∨_ : P n → P n → P n 36 | _`⇒_ : P n → P n → P n 37 | ?? : Fin n → P n 38 | 39 | unquoteDecl eqP = deriveEq eqP (quote P) 40 | 41 | Ctx : Nat → Set 42 | Ctx n = List (P n) 43 | 44 | data _⊢_ {n : Nat} (Γ : Ctx n) : P n → Set where 45 | var : ∀ {X} → X ∈ Γ → Γ ⊢ X 46 | lam : ∀ {X Y} → (X ∷ Γ) ⊢ Y → Γ ⊢ (X `⇒ Y) 47 | tt : Γ ⊢ `⊤ 48 | _,_ : ∀ {X Y} → Γ ⊢ X → Γ ⊢ Y → Γ ⊢ (X `∧ Y) 49 | inl : ∀ {X Y} → Γ ⊢ X → Γ ⊢ (X `∨ Y) 50 | inr : ∀ {X Y} → Γ ⊢ Y → Γ ⊢ (X `∨ Y) 51 | 52 | Env : Nat → Set₁ 53 | Env n = Vec Set n 54 | 55 | ⟦_⟧ : ∀ {n} → P n → Env n → Set 56 | ⟦ `⊤ ⟧ Δ = ⊤ 57 | ⟦ `⊥ ⟧ Δ = ⊥ 58 | ⟦ X `∧ Y ⟧ Δ = ⟦ X ⟧ Δ × ⟦ Y ⟧ Δ 59 | ⟦ X `∨ Y ⟧ Δ = Either (⟦ X ⟧ Δ) (⟦ Y ⟧ Δ) 60 | ⟦ X `⇒ Y ⟧ Δ = (⟦ X ⟧ Δ) → (⟦ Y ⟧ Δ) 61 | ⟦ ?? n ⟧ Δ = indexVec Δ n 62 | 63 | ⟦_⟧ctx : ∀ {n} → Ctx n → Env n → Set 64 | ⟦ Γ ⟧ctx Δ = All (λ A → ⟦ A ⟧ Δ) Γ 65 | 66 | 67 | ⊢-sound : ∀ {n} {Γ : Ctx n} {A : P n} → Γ ⊢ A → {Δ : Env n} → ⟦ Γ ⟧ctx Δ → ⟦ A ⟧ Δ 68 | ⊢-sound (var x) γ = lookup∈ γ x 69 | ⊢-sound tt γ = tt 70 | ⊢-sound (d₁ , d₂) γ = (⊢-sound d₁ γ) , (⊢-sound d₂ γ) 71 | ⊢-sound (inl d) γ = left (⊢-sound d γ) 72 | ⊢-sound (inr d) γ = right (⊢-sound d γ) 73 | ⊢-sound (lam d) γ z = ⊢-sound d (z ∷ γ) 74 | 75 | 76 | QEnv : Nat → Set 77 | QEnv n = Vec Type n 78 | 79 | record ParsedType (n : Nat) : Set where 80 | constructor parsed 81 | field 82 | {#env} : Nat 83 | env : QEnv #env 84 | lifter : P n → P #env 85 | type : P #env 86 | open ParsedType 87 | 88 | liftP : {n : Nat} (k : Nat) → P n → P (k + n) 89 | liftP k `⊤ = `⊤ 90 | liftP k `⊥ = `⊥ 91 | liftP k (A `∧ B) = liftP k A `∧ liftP k B 92 | liftP k (A `∨ B) = liftP k A `∨ liftP k B 93 | liftP k (A `⇒ B) = liftP k A `⇒ liftP k B 94 | liftP k ( ?? x ) = ?? (raise k x) 95 | 96 | {-# TERMINATING #-} 97 | parseP : Type → TC (ParsedType 0) 98 | parseP = loop [] 99 | 100 | where 101 | strengthn : Arg Type → Type → Term → TC (Maybe Term) 102 | strengthn dom a t = do 103 | t' ← newMeta a 104 | (extendContext dom (unify t (weaken 1 t')) >> return (just t')) 105 | <|> (return nothing) 106 | 107 | fallback : {n : Nat} → QEnv n → Type → TC (ParsedType n) 108 | fallback Δ t = case findIndex Δ t of λ where 109 | nothing → return (parsed (t ∷ Δ) (liftP 1) (?? 0)) 110 | (just x) → return (parsed Δ id (?? x)) 111 | 112 | loop : {n : Nat} → QEnv n → Type → TC (ParsedType n) 113 | loop Δ t = reduce t >>= λ where 114 | (def (quote ⊤) []) → return (parsed Δ id `⊤) 115 | (def (quote ⊥) []) → return (parsed Δ id `⊥) 116 | (def (quote Either) (_ ∷ _ ∷ (arg _ x) ∷ (arg _ y) ∷ [])) → do 117 | (parsed Δ₁ lift₁ X) ← loop Δ x 118 | (parsed Δ₂ lift₂ Y) ← loop Δ₁ y 119 | return (parsed Δ₂ (lift₂ ∘ lift₁) (lift₂ X `∨ Y)) 120 | (def (quote Σ) (_ ∷ _ ∷ ax@(arg _ x) ∷ (arg _ (lam _ (abs _ y))) ∷ [])) → do 121 | just y ← strengthn ax set! y 122 | where nothing → fallback Δ t 123 | (parsed Δ₁ lift₁ X) ← loop Δ x 124 | (parsed Δ₂ lift₂ Y) ← loop Δ₁ y 125 | return (parsed Δ₂ (lift₂ ∘ lift₁) (lift₂ X `∧ Y)) 126 | t@(pi ax@(arg _ x) (abs _ y)) → do 127 | just y ← strengthn ax set! y 128 | where nothing → fallback Δ t 129 | (parsed Δ₁ lift₁ X) ← loop Δ x 130 | (parsed Δ₂ lift₂ Y) ← loop Δ₁ y 131 | return (parsed Δ₂ (lift₂ ∘ lift₁) (lift₂ X `⇒ Y)) 132 | t → fallback Δ t 133 | 134 | macro 135 | parse : Term → Term → TC ⊤ 136 | parse t goal = do 137 | T ← parseP t 138 | result ← quoteTC T 139 | unify goal result 140 | 141 | testType₁ : Set 142 | testType₁ = ⊤ × ⊥ → Either ⊥ ⊤ 143 | 144 | testParse₁ : parse (testType₁) ≡ parsed [] _ ((`⊤ `∧ `⊥) `⇒ (`⊥ `∨ `⊤)) 145 | testParse₁ = refl 146 | 147 | testType₂ : Set → Set 148 | testType₂ A = A × ⊤ → Either ⊥ A 149 | 150 | testParse₂ : ∀ {A} → parse (testType₂ A) ≡ parsed (var₀ 0 ∷ []) _ ((?? zero `∧ `⊤) `⇒ (`⊥ `∨ ?? zero)) 151 | testParse₂ = refl 152 | 153 | assumption : ∀ {n} (Γ : Ctx n) (X : P n) → Maybe (X ∈ Γ) 154 | assumption [] Y = nothing 155 | assumption (X ∷ Γ) Y with X == Y 156 | assumption (X ∷ Γ) .X | yes refl = just zero! 157 | assumption (X ∷ Γ) Y | no x = suc <$> assumption Γ Y 158 | 159 | solveTauto : ∀ {n} (Γ : Ctx n) (X : P n) → Maybe (Γ ⊢ X) 160 | solveTauto {n} Γ X = (var <$> assumption Γ X) <|> solveTautoAux X 161 | where 162 | solveTautoAux : (X : P n) → Maybe (Γ ⊢ X) 163 | solveTautoAux = λ where 164 | `⊤ → return tt 165 | `⊥ → nothing 166 | (X `∧ Y) → _,_ <$> solveTauto Γ X <*> solveTauto Γ Y 167 | (X `∨ Y) → (inl <$> solveTauto Γ X) <|> (inr <$> solveTauto Γ Y) 168 | (X `⇒ Y) → lam <$> solveTauto (X ∷ Γ) Y 169 | (?? x ) → nothing 170 | 171 | macro 172 | tauto : Term → TC ⊤ 173 | tauto goal = do 174 | x ← inferType goal 175 | parsed Δ lift X ← parseP x 176 | case solveTauto [] X of λ where 177 | (just x) → do 178 | Δ ← traverse′ (unquoteTC {A = Set}) Δ 179 | let proof : ⟦ X ⟧ Δ 180 | proof = ⊢-sound x [] 181 | unify goal =<< quoteTC proof 182 | nothing → do 183 | X ← quoteTC X 184 | typeError (strErr "Failed to solve" ∷ termErr X ∷ []) 185 | 186 | 187 | test₁ : ⊤ 188 | test₁ = tauto 189 | 190 | test₂ : ⊥ → ⊥ 191 | test₂ = tauto 192 | 193 | test₃ : ⊥ → ⊤ → ⊥ × ⊤ 194 | test₃ = tauto 195 | 196 | test₄ : Either ⊥ ⊤ 197 | test₄ = tauto 198 | 199 | test₅ : ⊤ → Either ⊤ ⊥ 200 | test₅ = tauto 201 | 202 | module _ {A B C : Set} where 203 | 204 | test₇ : A → Either ⊥ A 205 | test₇ = tauto 206 | 207 | test₈ : A → B → A × B 208 | test₈ = tauto 209 | 210 | test₉ : A → (Either C (Either B A)) 211 | test₉ = tauto 212 | --------------------------------------------------------------------------------