module README.Tactic.RingSolver where
open import Agda.Builtin.FromNat
open import Data.Nat using (ℕ)
open import Data.Integer using (ℤ)
import Data.Nat.Literals as ℕ
import Data.Integer.Literals as ℤ
instance
numberNat : Number ℕ
numberNat = ℕ.number
instance
numberInt : Number ℤ
numberInt = ℤ.number
open import Data.List as List using (List; _∷_; [])
open import Function
open import Relation.Binary.PropositionalEquality as ≡
using (subst; _≡_; module ≡-Reasoning)
open import Data.Bool as Bool using (Bool; true; false; if_then_else_)
open import Data.Unit using (⊤; tt)
open import Tactic.RingSolver.Core.AlmostCommutativeRing using (AlmostCommutativeRing)
module IntegerExamples where
open import Data.Integer.Tactic.RingSolver
open AlmostCommutativeRing ring
lemma₁ : ∀ x y → x + y * 1 + 3 ≈ 3 + 1 + y + x + - 1
lemma₁ = solve-∀
lemma₂ : ∀ x y → (x + y) ^ 2 ≈ x ^ 2 + 2 * x * y + y ^ 2
lemma₂ = solve-∀
lemma₃ : ∀ x y → x + y * 1 + 3 ≈ 2 + 1 + y + x
lemma₃ x y = begin
x + y * 1 + 3 ≡⟨ +-comm x (y * 1) ⟨ +-cong ⟩ refl ⟩
y * 1 + x + 3 ≡⟨ solve (x ∷ y ∷ []) ⟩
3 + y + x ≡⟨⟩
2 + 1 + y + x ∎
where open ≡-Reasoning
module NaturalExamples where
open import Data.Nat.Tactic.RingSolver
open AlmostCommutativeRing ring
lemma₁ : ∀ x y → x + y * 1 + 3 ≈ 2 + 1 + y + x
lemma₁ = solve-∀
module _ {a} {A : Set a} (_≤_ : A → A → Bool) where
open import Data.Nat.Tactic.RingSolver
open AlmostCommutativeRing ring
data Tree : ℕ → Set a where
leaf : Tree 0
node : ∀ {n m} → A → Tree n → Tree m → Tree (1 + n + m)
infixr 1 _⇒_
_⇒_ : ∀ {n} → Tree n → ∀ {m} → n ≈ m → Tree m
x ⇒ n≈m = subst Tree n≈m x
open ≡-Reasoning
_∪_ : ∀ {n m} → Tree n → Tree m → Tree (n + m)
leaf ∪ ys = ys
node {a} {b} x xl xr ∪ leaf =
node x xl xr ⇒ solve (a ∷ b ∷ [])
node {a} {b} x xl xr ∪ node {c} {d} y yl yr =
if x ≤ y
then node x (node y yl yr ∪ xr) xl ⇒ begin
1 + (1 + c + d + b) + a ≡⟨ solve (a ∷ b ∷ c ∷ d ∷ []) ⟩
1 + a + b + (1 + c + d) ∎
else node y (node x xl xr ∪ yr) yl ⇒ begin
1 + (1 + a + b + d) + c ≡⟨ solve (a ∷ b ∷ c ∷ d ∷ []) ⟩
1 + a + b + (1 + c + d) ∎