diff --git a/src/Data/Integer/IntConstruction.agda b/src/Data/Integer/IntConstruction.agda new file mode 100644 index 0000000000..3e8e64389c --- /dev/null +++ b/src/Data/Integer/IntConstruction.agda @@ -0,0 +1,77 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Construction of integers as a pair of naturals +------------------------------------------------------------------------ + +{-# OPTIONS --safe --cubical-compatible #-} + +module Data.Integer.IntConstruction where + +open import Data.Nat.Base as ℕ using (ℕ) +open import Function.Base using (flip) +open import Relation.Binary +open import Relation.Binary.PropositionalEquality + +infixl 6 _⊖_ + +record ℤ : Set where + constructor _⊖_ + field + minuend : ℕ + subtrahend : ℕ + +infix 4 _≃_ _≤_ _≥_ _<_ _>_ + +record _≃_ (i j : ℤ) : Set where + constructor mk≃ + field + eq : ℤ.minuend i ℕ.+ ℤ.subtrahend j ≡ ℤ.minuend j ℕ.+ ℤ.subtrahend i + +_≤_ : Rel ℤ _ +(a ⊖ b) ≤ (c ⊖ d) = a ℕ.+ d ℕ.≤ c ℕ.+ b + +_≥_ : Rel ℤ _ +_≥_ = flip _≤_ + +_<_ : Rel ℤ _ +(a ⊖ b) < (c ⊖ d) = a ℕ.+ d ℕ.< c ℕ.+ b + +_>_ : Rel ℤ _ +_>_ = flip _<_ + +suc : ℤ → ℤ +suc (a ⊖ b) = ℕ.suc a ⊖ b + +pred : ℤ → ℤ +pred (a ⊖ b) = a ⊖ ℕ.suc b + +0ℤ : ℤ +0ℤ = 0 ⊖ 0 + +1ℤ : ℤ +1ℤ = 1 ⊖ 0 + +infixl 6 _+_ +_+_ : ℤ → ℤ → ℤ +(a ⊖ b) + (c ⊖ d) = (a ℕ.+ c) ⊖ (b ℕ.+ d) + +infixl 7 _*_ +_*_ : ℤ → ℤ → ℤ +(a ⊖ b) * (c ⊖ d) = (a ℕ.* c ℕ.+ b ℕ.* d) ⊖ (a ℕ.* d ℕ.+ b ℕ.* c) + +infix 8 -_ +-_ : ℤ → ℤ +- (a ⊖ b) = b ⊖ a + +infix 8 ⁻_ ⁺_ + +⁺_ : ℕ → ℤ +⁺ n = n ⊖ 0 + +⁻_ : ℕ → ℤ +⁻ n = 0 ⊖ n + +∣_∣ : ℤ → ℕ +∣ a ⊖ b ∣ = ℕ.∣ a - b ∣′ + diff --git a/src/Data/Integer/IntConstruction/DivMod.agda b/src/Data/Integer/IntConstruction/DivMod.agda new file mode 100644 index 0000000000..6d21c7f64a --- /dev/null +++ b/src/Data/Integer/IntConstruction/DivMod.agda @@ -0,0 +1,108 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Division with remainder for integers constructed as a pair of naturals +------------------------------------------------------------------------ + +{-# OPTIONS --safe --cubical-compatible #-} + +module Data.Integer.IntConstruction.DivMod where + +open import Data.Fin.Base as Fin using (Fin) +import Data.Fin.Properties as Fin +open import Data.Integer.IntConstruction +open import Data.Integer.IntConstruction.Properties +open import Data.Nat.Base as ℕ using (ℕ) +import Data.Nat.Properties as ℕ +import Data.Nat.DivMod as ℕ +open import Function.Base +open import Relation.Binary +open import Relation.Binary.PropositionalEquality +open import Relation.Nullary.Decidable + +-- should the divisor also be an integer? +record DivMod (dividend : ℤ) (divisor : ℕ) : Set where + field + quotient : ℤ + remainder : ℕ + remainder b%n≰a%n)) (ℕ.≤-trans (ℕ.m∸n≤m (b ℕ.% n) (a ℕ.% n)) (ℕ.<⇒≤ (ℕ.m%nj = ℕ.+-cancelʳ-< (d ℕ.+ c) (e ℕ.+ b) (a ℕ.+ f) $ trans-helper ℕ._>_ a b c d e f (ℕ.+-mono-<-≤ i>j (ℕ.≤-reflexive (sym j≃k))) + +<-respʳ-≃ : _<_ Respectsʳ _≃_ +<-respʳ-≃ {a ⊖ b} {c ⊖ d} {e ⊖ f} (mk≃ j≃k) i a+d≮c+b a+d≢c+d a+d>c+d = tri> a+d≮c+b (a+d≢c+d ∘ _≃_.eq) a+d>c+d + +------------------------------------------------------------------------ +-- Algebraic properties of _+_ + ++-cong : Congruent₂ _+_ ++-cong {a ⊖ b} {c ⊖ d} {e ⊖ f} {g ⊖ h} (mk≃ a+d≡c+b) (mk≃ e+h≡g+f) = mk≃ $ begin + (a ℕ.+ e) ℕ.+ (d ℕ.+ h) ≡⟨ medial a e d h ⟩ + (a ℕ.+ d) ℕ.+ (e ℕ.+ h) ≡⟨ cong₂ ℕ._+_ a+d≡c+b e+h≡g+f ⟩ + (c ℕ.+ b) ℕ.+ (g ℕ.+ f) ≡⟨ medial c b g f ⟩ + (c ℕ.+ g) ℕ.+ (b ℕ.+ f) ∎ + where + open ≡-Reasoning + open CommSemigroupProps ℕ.+-commutativeSemigroup + ++-assoc : Associative _+_ ++-assoc (a ⊖ b) (c ⊖ d) (e ⊖ f) = mk≃ $ cong₂ ℕ._+_ (ℕ.+-assoc a c e) (sym (ℕ.+-assoc b d f)) + ++-identityˡ : LeftIdentity 0ℤ _+_ ++-identityˡ (a ⊖ b) = mk≃ refl + ++-identityʳ : RightIdentity 0ℤ _+_ ++-identityʳ (a ⊖ b) = mk≃ $ cong₂ ℕ._+_ (ℕ.+-identityʳ a) (sym (ℕ.+-identityʳ b)) + ++-comm : Commutative _+_ ++-comm (a ⊖ b) (c ⊖ d) = mk≃ $ cong₂ ℕ._+_ (ℕ.+-comm a c) (ℕ.+-comm d b) + +------------------------------------------------------------------------ +-- Properties of _+_ and _≤_ + ++-mono-≤ : Monotonic₂ _≤_ _≤_ _≤_ _+_ ++-mono-≤ {a ⊖ b} {c ⊖ d} {e ⊖ f} {g ⊖ h} a+d≤c+b e+h≤g+f = begin + (a ℕ.+ e) ℕ.+ (d ℕ.+ h) ≡⟨ medial a e d h ⟩ + (a ℕ.+ d) ℕ.+ (e ℕ.+ h) ≤⟨ ℕ.+-mono-≤ a+d≤c+b e+h≤g+f ⟩ + (c ℕ.+ b) ℕ.+ (g ℕ.+ f) ≡⟨ medial c b g f ⟩ + (c ℕ.+ g) ℕ.+ (b ℕ.+ f) ∎ + where + open ℕ.≤-Reasoning + open CommSemigroupProps ℕ.+-commutativeSemigroup + +------------------------------------------------------------------------ +-- Algebraic properties of -_ + +-‿cong : Congruent₁ -_ +-‿cong {a ⊖ b} {c ⊖ d} (mk≃ a+d≡c+b) = mk≃ $ begin + b ℕ.+ c ≡⟨ ℕ.+-comm b c ⟩ + c ℕ.+ b ≡⟨ a+d≡c+b ⟨ + a ℕ.+ d ≡⟨ ℕ.+-comm a d ⟩ + d ℕ.+ a ∎ + where open ≡-Reasoning + ++-inverseˡ : LeftInverse 0ℤ -_ _+_ ++-inverseˡ (a ⊖ b) = mk≃ $ trans (ℕ.+-identityʳ (b ℕ.+ a)) (ℕ.+-comm b a) + ++-inverseʳ : RightInverse 0ℤ -_ _+_ ++-inverseʳ (a ⊖ b) = mk≃ $ trans (ℕ.+-identityʳ (a ℕ.+ b)) (ℕ.+-comm a b) + +------------------------------------------------------------------------ +-- Properties of -_ and _≤_ + +-‿anti-≤ : Antitonic₁ _≤_ _≤_ -_ +-‿anti-≤ {a ⊖ b} {c ⊖ d} a+d≥c+b = begin + b ℕ.+ c ≡⟨ ℕ.+-comm b c ⟩ + c ℕ.+ b ≤⟨ a+d≥c+b ⟩ + a ℕ.+ d ≡⟨ ℕ.+-comm a d ⟩ + d ℕ.+ a ∎ + where open ℕ.≤-Reasoning + +------------------------------------------------------------------------ +-- Algebraic properties of _*_ + +*-cong : Congruent₂ _*_ +*-cong {a ⊖ b} {c ⊖ d} {e ⊖ f} {g ⊖ h} (mk≃ a+d≡c+b) (mk≃ e+h≡g+f) = mk≃ $ ℕ.+-cancelʳ-≡ (d ℕ.* e) _ _ $ begin + ((a ℕ.* e ℕ.+ b ℕ.* f) ℕ.+ (c ℕ.* h ℕ.+ d ℕ.* g)) ℕ.+ d ℕ.* e ≡⟨ ℕ.+-assoc (a ℕ.* e ℕ.+ b ℕ.* f) (c ℕ.* h ℕ.+ d ℕ.* g) (d ℕ.* e) ⟩ + (a ℕ.* e ℕ.+ b ℕ.* f) ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ d ℕ.* e) ≡⟨ cong (a ℕ.* e ℕ.+ b ℕ.* f ℕ.+_) (ℕ.+-comm (c ℕ.* h ℕ.+ d ℕ.* g) (d ℕ.* e)) ⟩ + (a ℕ.* e ℕ.+ b ℕ.* f) ℕ.+ (d ℕ.* e ℕ.+ (c ℕ.* h ℕ.+ d ℕ.* g)) ≡⟨ medial (a ℕ.* e) (b ℕ.* f) (d ℕ.* e) (c ℕ.* h ℕ.+ d ℕ.* g) ⟩ + (a ℕ.* e ℕ.+ d ℕ.* e) ℕ.+ (b ℕ.* f ℕ.+ (c ℕ.* h ℕ.+ d ℕ.* g)) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ e a d) (ℕ.+-comm (c ℕ.* h ℕ.+ d ℕ.* g) (b ℕ.* f)) ⟨ + (a ℕ.+ d) ℕ.* e ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong (λ k → k ℕ.* e ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f)) a+d≡c+b ⟩ + (c ℕ.+ b) ℕ.* e ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ e c b) (ℕ.+-assoc (c ℕ.* h) (d ℕ.* g) (b ℕ.* f)) ⟩ + (c ℕ.* e ℕ.+ b ℕ.* e) ℕ.+ (c ℕ.* h ℕ.+ (d ℕ.* g ℕ.+ b ℕ.* f)) ≡⟨ medial (c ℕ.* e) (b ℕ.* e) (c ℕ.* h) (d ℕ.* g ℕ.+ b ℕ.* f) ⟩ + (c ℕ.* e ℕ.+ c ℕ.* h) ℕ.+ (b ℕ.* e ℕ.+ (d ℕ.* g ℕ.+ b ℕ.* f)) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribˡ-+ c e h) (ℕ.+-assoc (b ℕ.* e) (d ℕ.* g) (b ℕ.* f)) ⟨ + c ℕ.* (e ℕ.+ h) ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong (λ k → c ℕ.* k ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f)) e+h≡g+f ⟩ + c ℕ.* (g ℕ.+ f) ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong (ℕ._+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f)) (ℕ.*-distribˡ-+ c g f) ⟩ + (c ℕ.* g ℕ.+ c ℕ.* f) ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ medial (c ℕ.* g) (c ℕ.* f) (b ℕ.* e ℕ.+ d ℕ.* g) (b ℕ.* f) ⟩ + (c ℕ.* g ℕ.+ (b ℕ.* e ℕ.+ d ℕ.* g)) ℕ.+ (c ℕ.* f ℕ.+ b ℕ.* f) ≡⟨ cong₂ ℕ._+_ (ℕ.+-assoc (c ℕ.* g) (b ℕ.* e) (d ℕ.* g)) (ℕ.*-distribʳ-+ f c b) ⟨ + ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ (c ℕ.+ b) ℕ.* f ≡⟨ cong (λ k → ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ k ℕ.* f) a+d≡c+b ⟨ + ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ (a ℕ.+ d) ℕ.* f ≡⟨ cong (c ℕ.* g ℕ.+ b ℕ.* e ℕ.+ d ℕ.* g ℕ.+_) (ℕ.*-distribʳ-+ f a d) ⟩ + ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ (a ℕ.* f ℕ.+ d ℕ.* f) ≡⟨ medial (c ℕ.* g ℕ.+ b ℕ.* e) (d ℕ.* g) (a ℕ.* f) (d ℕ.* f) ⟩ + ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+ (d ℕ.* g ℕ.+ d ℕ.* f) ≡⟨ cong (((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+_) (ℕ.*-distribˡ-+ d g f) ⟨ + ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+ d ℕ.* (g ℕ.+ f) ≡⟨ cong (λ k → c ℕ.* g ℕ.+ b ℕ.* e ℕ.+ a ℕ.* f ℕ.+ d ℕ.* k) e+h≡g+f ⟨ + ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+ d ℕ.* (e ℕ.+ h) ≡⟨ cong₂ (λ j k → j ℕ.+ d ℕ.* k) (ℕ.+-assoc (c ℕ.* g) (b ℕ.* e) (a ℕ.* f)) (ℕ.+-comm e h) ⟩ + (c ℕ.* g ℕ.+ (b ℕ.* e ℕ.+ a ℕ.* f)) ℕ.+ d ℕ.* (h ℕ.+ e) ≡⟨ cong₂ (λ j k → c ℕ.* g ℕ.+ j ℕ.+ k) (ℕ.+-comm (b ℕ.* e) (a ℕ.* f)) (ℕ.*-distribˡ-+ d h e) ⟩ + (c ℕ.* g ℕ.+ (a ℕ.* f ℕ.+ b ℕ.* e)) ℕ.+ (d ℕ.* h ℕ.+ d ℕ.* e) ≡⟨ medial (c ℕ.* g) (a ℕ.* f ℕ.+ b ℕ.* e) (d ℕ.* h) (d ℕ.* e) ⟩ + (c ℕ.* g ℕ.+ d ℕ.* h) ℕ.+ ((a ℕ.* f ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* e) ≡⟨ ℕ.+-assoc (c ℕ.* g ℕ.+ d ℕ.* h) (a ℕ.* f ℕ.+ b ℕ.* e) (d ℕ.* e) ⟨ + ((c ℕ.* g ℕ.+ d ℕ.* h) ℕ.+ (a ℕ.* f ℕ.+ b ℕ.* e)) ℕ.+ d ℕ.* e ∎ + where + open ≡-Reasoning + open CommSemigroupProps ℕ.+-commutativeSemigroup + +*-assoc : Associative _*_ +*-assoc (a ⊖ b) (c ⊖ d) (e ⊖ f) = mk≃ $ cong₂ ℕ._+_ (lemma a b c d e f) (sym (lemma a b c d f e)) + where + open ≡-Reasoning + open CommSemigroupProps ℕ.+-commutativeSemigroup + lemma : ∀ u v w x y z → (u ℕ.* w ℕ.+ v ℕ.* x) ℕ.* y ℕ.+ (u ℕ.* x ℕ.+ v ℕ.* w) ℕ.* z + ≡ u ℕ.* (w ℕ.* y ℕ.+ x ℕ.* z) ℕ.+ v ℕ.* (w ℕ.* z ℕ.+ x ℕ.* y) + lemma u v w x y z = begin + (u ℕ.* w ℕ.+ v ℕ.* x) ℕ.* y ℕ.+ (u ℕ.* x ℕ.+ v ℕ.* w) ℕ.* z ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ y (u ℕ.* w) (v ℕ.* x)) (ℕ.*-distribʳ-+ z (u ℕ.* x) (v ℕ.* w)) ⟩ + ((u ℕ.* w) ℕ.* y ℕ.+ (v ℕ.* x) ℕ.* y) ℕ.+ ((u ℕ.* x) ℕ.* z ℕ.+ (v ℕ.* w) ℕ.* z) ≡⟨ cong₂ ℕ._+_ (cong₂ ℕ._+_ (ℕ.*-assoc u w y) (ℕ.*-assoc v x y)) (cong₂ ℕ._+_ (ℕ.*-assoc u x z) (ℕ.*-assoc v w z)) ⟩ + (u ℕ.* (w ℕ.* y) ℕ.+ v ℕ.* (x ℕ.* y)) ℕ.+ (u ℕ.* (x ℕ.* z) ℕ.+ v ℕ.* (w ℕ.* z)) ≡⟨ medial (u ℕ.* (w ℕ.* y)) (v ℕ.* (x ℕ.* y)) (u ℕ.* (x ℕ.* z)) (v ℕ.* (w ℕ.* z)) ⟩ + (u ℕ.* (w ℕ.* y) ℕ.+ u ℕ.* (x ℕ.* z)) ℕ.+ (v ℕ.* (x ℕ.* y) ℕ.+ v ℕ.* (w ℕ.* z)) ≡⟨ cong ((u ℕ.* (w ℕ.* y) ℕ.+ u ℕ.* (x ℕ.* z)) ℕ.+_) (ℕ.+-comm (v ℕ.* (x ℕ.* y)) (v ℕ.* (w ℕ.* z))) ⟩ + (u ℕ.* (w ℕ.* y) ℕ.+ u ℕ.* (x ℕ.* z)) ℕ.+ (v ℕ.* (w ℕ.* z) ℕ.+ v ℕ.* (x ℕ.* y)) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribˡ-+ u (w ℕ.* y) (x ℕ.* z)) (ℕ.*-distribˡ-+ v (w ℕ.* z) (x ℕ.* y)) ⟨ + u ℕ.* (w ℕ.* y ℕ.+ x ℕ.* z) ℕ.+ v ℕ.* (w ℕ.* z ℕ.+ x ℕ.* y) ∎ + +*-zeroˡ : LeftZero 0ℤ _*_ +*-zeroˡ _ = mk≃ refl + +*-zeroʳ : RightZero 0ℤ _*_ +*-zeroʳ _ = mk≃ $ ℕ.+-identityʳ _ + +*-identityˡ : LeftIdentity 1ℤ _*_ +*-identityˡ (a ⊖ b) = mk≃ $ cong₂ ℕ._+_ (lemma a) (sym (lemma b)) + where + lemma : ∀ n → n ℕ.+ 0 ℕ.+ 0 ≡ n + lemma n = trans (ℕ.+-identityʳ (n ℕ.+ 0)) (ℕ.+-identityʳ n) + +*-identityʳ : RightIdentity 1ℤ _*_ +*-identityʳ (a ⊖ b) = mk≃ $ cong₂ ℕ._+_ l (sym r) + where + l : a ℕ.* 1 ℕ.+ b ℕ.* 0 ≡ a + l = trans (cong₂ ℕ._+_ (ℕ.*-identityʳ a) (ℕ.*-zeroʳ b)) (ℕ.+-identityʳ a) + r : a ℕ.* 0 ℕ.+ b ℕ.* 1 ≡ b + r = trans (cong₂ ℕ._+_ (ℕ.*-zeroʳ a) (ℕ.*-identityʳ b)) (ℕ.+-identityˡ b) + +*-distribˡ-+ : _*_ DistributesOverˡ _+_ +*-distribˡ-+ (a ⊖ b) (c ⊖ d) (e ⊖ f) = mk≃ $ cong₂ ℕ._+_ (lemma a b c d e f) (sym (lemma a b d c f e)) + where + open ≡-Reasoning + open CommSemigroupProps ℕ.+-commutativeSemigroup + lemma : ∀ u v w x y z → u ℕ.* (w ℕ.+ y) ℕ.+ v ℕ.* (x ℕ.+ z) + ≡ (u ℕ.* w ℕ.+ v ℕ.* x) ℕ.+ (u ℕ.* y ℕ.+ v ℕ.* z) + lemma u v w x y z = begin + u ℕ.* (w ℕ.+ y) ℕ.+ v ℕ.* (x ℕ.+ z) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribˡ-+ u w y) (ℕ.*-distribˡ-+ v x z) ⟩ + (u ℕ.* w ℕ.+ u ℕ.* y) ℕ.+ (v ℕ.* x ℕ.+ v ℕ.* z) ≡⟨ medial (u ℕ.* w) (u ℕ.* y) (v ℕ.* x) (v ℕ.* z) ⟩ + (u ℕ.* w ℕ.+ v ℕ.* x) ℕ.+ (u ℕ.* y ℕ.+ v ℕ.* z) ∎ + +*-distribʳ-+ : _*_ DistributesOverʳ _+_ +*-distribʳ-+ (a ⊖ b) (c ⊖ d) (e ⊖ f) = mk≃ $ cong₂ ℕ._+_ (lemma a b c d e f) (sym (lemma b a c d e f)) + where + open ≡-Reasoning + open CommSemigroupProps ℕ.+-commutativeSemigroup + lemma : ∀ u v w x y z → (w ℕ.+ y) ℕ.* u ℕ.+ (x ℕ.+ z) ℕ.* v + ≡ (w ℕ.* u ℕ.+ x ℕ.* v) ℕ.+ (y ℕ.* u ℕ.+ z ℕ.* v) + lemma u v w x y z = begin + (w ℕ.+ y) ℕ.* u ℕ.+ (x ℕ.+ z) ℕ.* v ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ u w y) (ℕ.*-distribʳ-+ v x z) ⟩ + (w ℕ.* u ℕ.+ y ℕ.* u) ℕ.+ (x ℕ.* v ℕ.+ z ℕ.* v) ≡⟨ medial (w ℕ.* u) (y ℕ.* u) (x ℕ.* v) (z ℕ.* v) ⟩ + (w ℕ.* u ℕ.+ x ℕ.* v) ℕ.+ (y ℕ.* u ℕ.+ z ℕ.* v) ∎ + +*-comm : Commutative _*_ +*-comm (a ⊖ b) (c ⊖ d) = mk≃ $ cong₂ ℕ._+_ (cong₂ ℕ._+_ (ℕ.*-comm a c) (ℕ.*-comm b d)) (trans (ℕ.+-comm (c ℕ.* b) (d ℕ.* a)) (cong₂ ℕ._+_ (ℕ.*-comm d a) (ℕ.*-comm c b))) + +------------------------------------------------------------------------ +-- Properties of ⁺_ + +⁺-cong : ∀ {m n} → m ≡ n → ⁺ m ≃ ⁺ n +⁺-cong refl = mk≃ refl + +⁺-injective : ∀ {m n} → ⁺ m ≃ ⁺ n → m ≡ n +⁺-injective {m} {n} (mk≃ m+0≡n+0) = begin + m ≡⟨ ℕ.+-identityʳ m ⟨ + m ℕ.+ 0 ≡⟨ m+0≡n+0 ⟩ + n ℕ.+ 0 ≡⟨ ℕ.+-identityʳ n ⟩ + n ∎ + where open ≡-Reasoning + +⁺-mono-≤ : Monotonic₁ ℕ._≤_ _≤_ ⁺_ +⁺-mono-≤ = ℕ.+-monoˡ-≤ 0 + +⁺-mono-< : Monotonic₁ ℕ._<_ _<_ ⁺_ +⁺-mono-< = ℕ.+-monoˡ-< 0 + +⁺-+-homo : ∀ m n → ⁺ (m ℕ.+ n) ≃ ⁺ m + ⁺ n +⁺-+-homo m n = mk≃ refl + +⁺-0-homo : ⁺ 0 ≡ 0ℤ +⁺-0-homo = refl + +⁺-*-homo : ∀ m n → ⁺ (m ℕ.* n) ≃ ⁺ m * ⁺ n +⁺-*-homo m n = mk≃ $ begin + m ℕ.* n ℕ.+ (m ℕ.* 0 ℕ.+ 0) ≡⟨ cong (m ℕ.* n ℕ.+_) (ℕ.+-identityʳ (m ℕ.* 0)) ⟩ + m ℕ.* n ℕ.+ m ℕ.* 0 ≡⟨ cong (m ℕ.* n ℕ.+_) (ℕ.*-zeroʳ m) ⟩ + m ℕ.* n ℕ.+ 0 ≡⟨ ℕ.+-identityʳ (m ℕ.* n ℕ.+ 0) ⟨ + m ℕ.* n ℕ.+ 0 ℕ.+ 0 ∎ + where open ≡-Reasoning + +------------------------------------------------------------------------ +-- Properties of ∣_∣ + +∣∣-cong : ∀ {i j} → i ≃ j → ∣ i ∣ ≡ ∣ j ∣ +∣∣-cong {a ⊖ b} {c ⊖ d} (mk≃ a+d≡c+b) = begin + ℕ.∣ a - b ∣′ ≡⟨ ℕ.∣-∣≡∣-∣′ a b ⟨ + ℕ.∣ a - b ∣ ≡⟨ ℕ.∣m+o-n+o∣≡∣m-n∣ a b d ⟨ + ℕ.∣ a ℕ.+ d - b ℕ.+ d ∣ ≡⟨ cong ℕ.∣_- b ℕ.+ d ∣ a+d≡c+b ⟩ + ℕ.∣ c ℕ.+ b - b ℕ.+ d ∣ ≡⟨ cong ℕ.∣_- b ℕ.+ d ∣ (ℕ.+-comm c b) ⟩ + ℕ.∣ b ℕ.+ c - b ℕ.+ d ∣ ≡⟨ ℕ.∣m+n-m+o∣≡∣n-o∣ b c d ⟩ + ℕ.∣ c - d ∣ ≡⟨ ℕ.∣-∣≡∣-∣′ c d ⟩ + ℕ.∣ c - d ∣′ ∎ + where open ≡-Reasoning + +∣⁺_∣ : ∀ n → ∣ ⁺ n ∣ ≡ n +∣⁺ n ∣ = refl + +∣⁻_∣ : ∀ n → ∣ ⁻ n ∣ ≡ n +∣⁻ ℕ.zero ∣ = refl +∣⁻ ℕ.suc n ∣ = refl + +∣∣-surjective : ∀ n → ∃[ i ] ∀ {j} → j ≃ i → ∣ j ∣ ≡ n +∣∣-surjective n = ⁺ n , λ {j} j≃⁺n → trans (∣∣-cong {i = j} j≃⁺n) ∣⁺ n ∣ diff --git a/src/Data/Integer/IntConstruction/Tmp.agda b/src/Data/Integer/IntConstruction/Tmp.agda new file mode 100644 index 0000000000..a585346618 --- /dev/null +++ b/src/Data/Integer/IntConstruction/Tmp.agda @@ -0,0 +1,24 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- To be merged into Data.Integer.Base before merging! +------------------------------------------------------------------------ + +{-# OPTIONS --safe --cubical-compatible #-} + +module Data.Integer.IntConstruction.Tmp where + +open import Data.Integer.Base +import Data.Integer.IntConstruction as INT +import Data.Nat.Base as ℕ + +fromINT : INT.ℤ → ℤ +fromINT (m INT.⊖ n) = m ⊖ n + +toINT : ℤ → INT.ℤ +toINT (+ n) = n INT.⊖ 0 +toINT (-[1+ n ]) = 0 INT.⊖ ℕ.suc n + +-- here we have a choice: +-- * either we can show the definitions in Data.Int.Base match those on INT under the above isomorphism; +-- * or we could _redefine_ those operations _as_ the operations on INT under the above isormophism. diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda index dc40f03132..9c5dd8ebc5 100644 --- a/src/Data/Nat/Properties.agda +++ b/src/Data/Nat/Properties.agda @@ -1834,6 +1834,10 @@ m≤n⇒∣m-n∣≡n∸m {n = _} (s≤s m≤n) = m≤n⇒∣m-n∣≡n∸m ∣m+n-m+o∣≡∣n-o∣ zero n o = refl ∣m+n-m+o∣≡∣n-o∣ (suc m) n o = ∣m+n-m+o∣≡∣n-o∣ m n o +∣m+o-n+o∣≡∣m-n∣ : ∀ m n o → ∣ m + o - n + o ∣ ≡ ∣ m - n ∣ +∣m+o-n+o∣≡∣m-n∣ m n zero = cong₂ ∣_-_∣ (+-identityʳ m) (+-identityʳ n) +∣m+o-n+o∣≡∣m-n∣ m n (suc o) = trans (cong₂ ∣_-_∣ (+-suc m o) (+-suc n o)) (∣m+o-n+o∣≡∣m-n∣ m n o) + m∸n≤∣m-n∣ : ∀ m n → m ∸ n ≤ ∣ m - n ∣ m∸n≤∣m-n∣ m n with ≤-total m n ... | inj₁ m≤n = subst (_≤ ∣ m - n ∣) (sym (m≤n⇒m∸n≡0 m≤n)) z≤n