aboutsummaryrefslogtreecommitdiff
path: root/Prelude/Number.agda
blob: 02024d0e879633be50f7183b32908f702cc2d347 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
{- Copyright © 2015 Benjamin Barenblat

Licensed under the Apache License, Version 2.0 (the ‘License’); you may not use
this file except in compliance with the License.  You may obtain a copy of the
License at

    http://www.apache.org/licenses/LICENSE-2.0

Unless required by applicable law or agreed to in writing, software distributed
under the License is distributed on an ‘AS IS’ BASIS, WITHOUT WARRANTIES OR
CONDITIONS OF ANY KIND, either express or implied.  See the License for the
specific language governing permissions and limitations under the License. -}

module B.Prelude.Number where

import Algebra
import Data.Integer
import Data.Integer.Properties
import Data.Nat
import Data.Nat.Properties
open import Level using (_⊔_)
open import Relation.Binary.PropositionalEquality using (_≡_; refl)

open Data.Integer using (ℤ) public
open Data.Nat using (ℕ) public

record NearSemiring {c} {ℓ} (A : Set c) : Set (Level.suc (c ⊔ ℓ)) where
  field
    structure : Algebra.NearSemiring c ℓ
    typeEquality : A ≡ Algebra.NearSemiring.Carrier structure

NearSemiringInstance : ∀ {c ℓ}
                     → (structure : Algebra.NearSemiring c ℓ)
                     → NearSemiring (Algebra.NearSemiring.Carrier structure)
NearSemiringInstance structure =
  record { structure = structure; typeEquality = refl }

private
  change₂ : ∀ {c ℓ A′}
          → ⦃ witness : NearSemiring {c} {ℓ} A′ ⦄
          → ((r : Algebra.NearSemiring c ℓ)
             → let A = Algebra.NearSemiring.Carrier r in
               A → A → A)
          → A′ → A′ → A′
  change₂ ⦃ witness ⦄ f rewrite NearSemiring.typeEquality witness =
    f (NearSemiring.structure witness)

_*_ : ∀ {c ℓ A} → ⦃ _ : NearSemiring {c} {ℓ} A ⦄ → A → A → A
_*_ = change₂ Algebra.NearSemiring._*_

instance
  NearSemiring-ℕ : NearSemiring ℕ
  NearSemiring-ℕ =
    NearSemiringInstance (Algebra.CommutativeSemiring.nearSemiring
                            Data.Nat.Properties.commutativeSemiring)

  NearSemiring-ℤ : NearSemiring ℤ
  NearSemiring-ℤ =
    NearSemiringInstance (Algebra.CommutativeRing.nearSemiring
                            Data.Integer.Properties.commutativeRing)