From 662ca19ed7dabecdf17cdd555132f106bb768e58 Mon Sep 17 00:00:00 2001 From: Greg Brown Date: Thu, 7 Apr 2022 17:58:39 +0100 Subject: Add a new ring solver tactic. --- .../Tactic/CommutativeRing/NonReflective.agda | 283 +++++++++++++++++++++ 1 file changed, 283 insertions(+) create mode 100644 src/Helium/Tactic/CommutativeRing/NonReflective.agda (limited to 'src/Helium/Tactic/CommutativeRing/NonReflective.agda') diff --git a/src/Helium/Tactic/CommutativeRing/NonReflective.agda b/src/Helium/Tactic/CommutativeRing/NonReflective.agda new file mode 100644 index 0000000..12029a5 --- /dev/null +++ b/src/Helium/Tactic/CommutativeRing/NonReflective.agda @@ -0,0 +1,283 @@ +-------------------------------------------------------------------------------- +-- Agda Helium +-- +-- Ring solver tactic using integers as coefficients +-------------------------------------------------------------------------------- + +{-# OPTIONS --without-K --safe #-} + +open import Helium.Algebra.Ordered.StrictTotal.Bundles using (CommutativeRing) +import Helium.Algebra.Ordered.StrictTotal.Properties.CommutativeRing as Ringₚ +open import Relation.Nullary + +module Helium.Tactic.CommutativeRing.NonReflective + {ℓ₁ ℓ₂ ℓ₃} + (R : CommutativeRing ℓ₁ ℓ₂ ℓ₃) + (1≉0 : let module R = CommutativeRing R in ¬ R.1# R.≈ R.0#) + where + +open import Data.Bool.Base +open import Data.Sign using (Sign) renaming (_*_ to _S*_) +open import Level using (0ℓ; _⊔_) + +open import Relation.Binary.Reasoning.StrictPartialOrder (CommutativeRing.strictPartialOrder R) + +module R where + open CommutativeRing R public + renaming + ( trans to <-trans + ; irrefl to <-irrefl + ; asym to <-asym + ; 00∧y>0⇒xy>0 + ) + + open Ringₚ R public + + infixr 2 _,_ + + record Positive : Set (ℓ₁ ⊔ ℓ₂ ⊔ ℓ₃) where + constructor _,_ + field + val : Carrier + proof : val ≥ 0# + + open Positive + + infixl 6 _*⁺_ _*′_ + infix 5 _◃_ + + _*⁺_ : Positive → Positive → Positive + +x *⁺ +y = +x .val * +y .val , x≥0∧y≥0⇒x*y≥0 (+x .proof) (+y .proof) + + _◃_ : Sign → Positive → Carrier + Sign.- ◃ +x = - (+x .val) + Sign.+ ◃ +x = +x .val + + sign : Carrier → Sign + sign x with does (x 0 x<0) + ... | no x≮0 = x , ≮⇒≥ x≮0 + + ◃-congˡ : ∀ s +x +y → +x .val ≈ +y .val → s ◃ +x ≈ s ◃ +y + ◃-congˡ Sign.- +x +y x≈y = -‿cong x≈y + ◃-congˡ Sign.+ +x +y x≈y = x≈y + + _*′_ : Carrier → Carrier → Carrier + x *′ y = sign x S* sign y ◃ ∣ x ∣ *⁺ ∣ y ∣ + + *′≈* : ∀ x y → x *′ y ≈ x * y + *′≈* x y with x 0⇒-x<0 (R.n≢0∧x>0⇒n×x>0 (suc n) (R.x≉y⇒0<1 1≉0)))) + + sign-homo : ∀ n → sign n ≡ R.sign ⟦ n ⟧ + sign-homo n with ⟦ n ⟧ R.0⇒-x<0 (R.n≢0∧x>0⇒n×x>0 (suc n) (R.x≉y⇒0<1 1≉0)))) + + ◃-homo : ∀ s n → ⟦ s ◃ n ⟧ R.≈ s R.◃ (⟦ + n ⟧ R., R.x≥0⇒n×x≥0 n R.0≤1) + ◃-homo Sign.- 0 = R.Eq.sym R.-0≈0 + ◃-homo Sign.- (suc n) = R.Eq.refl + ◃-homo Sign.+ 0 = R.Eq.refl + ◃-homo Sign.+ (suc n) = R.Eq.refl + + *ⁿ-homo : ∀ m n → ⟦ + (m ℕ.* n) ⟧ R.≈ ⟦ + m ⟧ R.* ⟦ + n ⟧ + *ⁿ-homo 0 n = R.Eq.sym (R.zeroˡ ⟦ + n ⟧) + *ⁿ-homo (suc m) n = begin-equality + (n ℕ.+ m ℕ.* n) R.× R.1# ≈⟨ R.×-homo-+ R.1# n (m ℕ.* n) ⟩ + n R.× R.1# R.+ (m ℕ.* n) R.× R.1# ≈⟨ R.+-congˡ (*ⁿ-homo m n) ⟩ + n R.× R.1# R.+ m R.× R.1# R.* n R.× R.1# ≈˘⟨ R.+-congʳ (R.*-identityˡ _) ⟩ + R.1# R.* n R.× R.1# R.+ m R.× R.1# R.* n R.× R.1# ≈˘⟨ R.distribʳ _ R.1# _ ⟩ + (R.1# R.+ m R.× R.1#) R.* n R.× R.1# ≈˘⟨ R.*-congʳ (R.1+× m R.1#) ⟩ + suc m R.× R.1# R.* n R.× R.1# ∎ + + *-homo : ∀ m n → ⟦ m * n ⟧ R.≈ ⟦ m ⟧ R.* ⟦ n ⟧ + *-homo m n = begin-equality + ⟦ m * n ⟧ ≡⟨⟩ + ⟦ sm*sn ◃ ∣ m ∣ ℕ.* ∣ n ∣ ⟧ ≈⟨ ◃-homo sm*sn (∣ m ∣ ℕ.* ∣ n ∣) ⟩ + sm*sn R.◃ (⟦ + (∣ m ∣ ℕ.* ∣ n ∣) ⟧ R., ∣m∣∣n∣≥0) ≈⟨ R.◃-congˡ sm*sn _ _ (*ⁿ-homo ∣ m ∣ ∣ n ∣) ⟩ + sm*sn R.◃ (⟦ + ∣ m ∣ ⟧ R.* ⟦ + ∣ n ∣ ⟧ R., ∣mn∣≥0) ≈⟨ R.◃-congˡ sm*sn _ _ (R.*-cong (∣∙∣-homo m) (∣∙∣-homo n)) ⟩ + sm*sn R.◃ R.∣ ⟦ m ⟧ ∣ R.*⁺ R.∣ ⟦ n ⟧ ∣ ≡⟨ cong₂ (λ x y → x S* y R.◃ R.∣ ⟦ m ⟧ ∣ R.*⁺ R.∣ ⟦ n ⟧ ∣) (sign-homo m) (sign-homo n) ⟩ + Rsm*Rsn R.◃ R.∣ ⟦ m ⟧ ∣ R.*⁺ R.∣ ⟦ n ⟧ ∣ ≡⟨⟩ + ⟦ m ⟧ R.*′ ⟦ n ⟧ ≈⟨ R.*′≈* ⟦ m ⟧ ⟦ n ⟧ ⟩ + ⟦ m ⟧ R.* ⟦ n ⟧ ∎ + where + ∣m∣∣n∣≥0 = R.x≥0⇒n×x≥0 (∣ m ∣ ℕ.* ∣ n ∣) R.0≤1 + ∣mn∣≥0 = R.x≥0∧y≥0⇒x*y≥0 (R.x≥0⇒n×x≥0 ∣ m ∣ R.0≤1) (R.x≥0⇒n×x≥0 ∣ n ∣ R.0≤1) + sm*sn = sign m S* sign n + Rsm*Rsn = R.sign ⟦ m ⟧ S* R.sign ⟦ n ⟧ + + homo : Homomorphism 0ℓ 0ℓ ℓ₁ ℓ₂ + homo = record + { from = ℤ-coeff + ; to = ring + ; morphism = ℤ⟶R + ; Zero-C⟶Zero-R = λ { +0 _ → R.Eq.refl } + } + + open Eval R.Unordered.rawRing (_-Raw-AlmostCommutative⟶_.⟦_⟧ ℤ⟶R) public + + open import Tactic.RingSolver.Core.Polynomial.Base (Homomorphism.from homo) + + norm : ∀ {n} → Expr ℤ n → Poly n + norm (Κ x) = κ x + norm (Ι x) = ι x + norm (x ⊕ y) = norm x ⊞ norm y + norm (x ⊗ y) = norm x ⊠ norm y + norm (x ⊛ i) = norm x ⊡ i + norm (⊝ x) = ⊟ norm x + + ⟦_⇓⟧ : ∀ {n} → Expr ℤ n → Vec R.Carrier n → R.Carrier + ⟦ expr ⇓⟧ = ⟦ norm expr ⟧ₚ + where open import Tactic.RingSolver.Core.Polynomial.Semantics homo renaming (⟦_⟧ to ⟦_⟧ₚ) + + correct : ∀ {n} (expr : Expr ℤ n) ρ → ⟦ expr ⇓⟧ ρ R.≈ ⟦ expr ⟧ ρ + correct {n = n} = go + where + open import Tactic.RingSolver.Core.Polynomial.Homomorphism homo + + go : ∀ (expr : Expr ℤ n) ρ → ⟦ expr ⇓⟧ ρ R.≈ ⟦ expr ⟧ ρ + go (Κ x) ρ = κ-hom x ρ + go (Ι x) ρ = ι-hom x ρ + go (x ⊕ y) ρ = ⊞-hom (norm x) (norm y) ρ ⟨ R.Eq.trans ⟩ (go x ρ ⟨ R.+-cong ⟩ go y ρ) + go (x ⊗ y) ρ = ⊠-hom (norm x) (norm y) ρ ⟨ R.Eq.trans ⟩ (go x ρ ⟨ R.*-cong ⟩ go y ρ) + go (x ⊛ i) ρ = ⊡-hom (norm x) i ρ ⟨ R.Eq.trans ⟩ R.^-congˡ i (go x ρ) + go (⊝ x) ρ = ⊟-hom (norm x) ρ ⟨ R.Eq.trans ⟩ R.-‿cong (go x ρ) + + open import Relation.Binary.Reflection R.strictTotalOrder.Eq.setoid Ι ⟦_⟧ ⟦_⇓⟧ correct public + +solve : ∀ (n : ℕ) → + (f : N-ary n (Expr ℤ n) (Expr ℤ n × Expr ℤ n)) → + Eqʰ n R._≈_ (curryⁿ (Ops.⟦_⇓⟧ (proj₁ (Ops.close n f)))) (curryⁿ (Ops.⟦_⇓⟧ (proj₂ (Ops.close n f)))) → + Eq n R._≈_ (curryⁿ (Ops.⟦_⟧ (proj₁ (Ops.close n f)))) (curryⁿ (Ops.⟦_⟧ (proj₂ (Ops.close n f)))) +solve = Ops.solve +{-# INLINE solve #-} + +infixl 6 _⊜_ +_⊜_ : ∀ {n} → Expr ℤ n → Expr ℤ n → Expr ℤ n × Expr ℤ n +_⊜_ = _,_ +{-# INLINE _⊜_ #-} -- cgit v1.2.3