├── .gitignore ├── README.md ├── agda-metric-reals.agda-lib ├── plan.org ├── shell.nix └── src ├── Algebra ├── MonoidOfSemigroup.agda └── Solver │ └── CommutativeSemigroup.agda ├── Data ├── Rational │ └── Unnormalised │ │ └── Positive.agda └── Real │ └── UpperClosed.agda ├── MetricSpace.agda └── MetricSpace ├── CartesianProduct.agda ├── Category.agda ├── CauchyReals.agda ├── Completion.agda ├── ConvexAlgebras.agda ├── Discrete.agda ├── InternalHom.agda ├── MonoidalProduct.agda ├── Rationals.agda ├── Scaling.agda └── Terminal.agda /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | *~ 3 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # agda-metric-reals 2 | 3 | This is all work in progress, but here's the table of contents / plan / aspirations: 4 | 5 | - "Upper" real numbers, defined as closed upper subsets of positive rationals [upper-reals](upper-reals.agda) 6 | - Each number is represented similarly to an upper Dedekind cut 7 | - These form a semiring and a complete lattice 8 | - Metric spaces, with distances valued in the upper reals [metric2/base](metric2/base.agda) 9 | - Form a category with non-expansive maps as morphisms 10 | - This category has a terminal object [metric2/terminal](metric2/terminal.agda). 11 | - It has binary products [metric2/product](metric2/product.agda), with distances given by maximum. 12 | - It has another monoidal structure [metric2/monoidal](metric2/monoidal.agda), with distance given by addition, which is closed [metric2/internal-hom](metric2/internal-hom.agda). 13 | - It has a graded "scaling" comonad, which allows the expression of all Lipschitz continuous functions by making their Lipschitz constant explicit [metric2/scaling](metric2/scaling.agda). 14 | - It has a completion monad, which completes a metric space by adding all limit points [metric2/completion](metric2/completion.agda). This is implemented using regular functions as in Russell O'Connor's work (see references below). 15 | - The rationals form a metric space [metric2/rationals](metric2/rationals.agda), whose completion is the (full) real numbers. This is mostly unfinished. 16 | - It has a monad of convex combinations [metric2/convex-alg](metric2/convex-alg.agda), i.e. a probability distribution monad. The basic idea follows the Quantitative Algebraic Theories of Mardare, Panangaden and Plotkin. This should allow an integration operator, following the work of O'Connor and Spitters. 17 | 18 | The Agda files are being developed with Agda 2.6.1 and standard library 1.6. 19 | 20 | ## References 21 | 22 | - Russell O'Connor: [A monadic, functional implementation of real numbers](https://arxiv.org/abs/cs/0605058). Math. Struct. Comput. Sci. 17(1): 129-159 (2007) 23 | - Russell O'Connor, Bas Spitters: [A computer-verified monadic functional implementation of the integral](https://doi.org/10.1016/j.tcs.2010.05.031). Theor. Comput. Sci. 411(37): 3386-3402 (2010) 24 | - Radu Mardare, Prakash Panangaden, Gordon D. Plotkin: [Quantitative Algebraic Reasoning](https://homepages.inf.ed.ac.uk/gdp/publications/Quantitative_Alg_Reasoning.pdf). LICS 2016: 700-709 25 | -------------------------------------------------------------------------------- /agda-metric-reals.agda-lib: -------------------------------------------------------------------------------- 1 | name: adga-metric-reals 2 | include: src 3 | depend: standard-library 4 | -------------------------------------------------------------------------------- /plan.org: -------------------------------------------------------------------------------- 1 | #+STARTUP: indent 2 | 3 | * Plan [0/3] 4 | - [ ] Rearrange files into a stdlib-like hierarchy 5 | - [ ] qpos -> Data.Rational.Positive ??? 6 | - [ ] CommutativeSemigroupSolver -> Algebra. 7 | - [ ] upper-reals -> Data.Real.UpperExtendee 8 | - [ ] Use Function.Metric.Rational 9 | - [ ] Finish bounding of rationals so can do multiplication 10 | - [ ] Maximum and minimum of rationals as non-expansive functions 11 | - [ ] Define ordering on rationals and reals using max 12 | - [ ] interface with agda-categories 13 | 14 | 15 | * Related work 16 | - https://github.com/andrejbauer/clerical/ 17 | - Floating-point arithmetic in the Coq system, Guillaume Melquiond 18 | - A Sheaf Theoretic Approach to Measure Theory, Matthew Jackson 19 | (PhD thesis). Uses the upper semicontinuous reals. 20 | - https://link.springer.com/book/10.1007/BFb0080643 -- Principles 21 | of Intuitionism 22 | - https://ncatlab.org/nlab/show/one-sided+real+number -- page about 23 | semicontinuous reals, listing some of the results 24 | - https://blogs.ed.ac.uk/he-lab/2021/03/18/todd-ambridge-global-optimisation-via-constructive-real-numbers/ 25 | - https://github.com/MrChico/Reals-in-agda 26 | 27 | ** A Universal Characterization of the Closed Euclidean Interval 28 | https://www.cs.bham.ac.uk/~mhe/papers/lics2001-revised.pdf 29 | 30 | - Martín H. Escardó, Alex K. Simpson 31 | - We propose a notion of interval object in a category with finite 32 | products, providing a universal property for closed and bounded real 33 | line segments. The universal property gives rise to an analogue of 34 | primitive recursion for defining computable functions on the 35 | interval. We use this to define basic arithmetic operations and to 36 | verify equations between them. We test the notion in categories of 37 | interest. In the category of sets, any closed and bounded interval 38 | of real numbers is an interval object. In the category of topo 39 | logical spaces, the interval objects are closed and bounded 40 | intervals with the Euclidean topology. We also prove that an 41 | interval object exists in any elementary topos with natural numbers 42 | object. 43 | 44 | * Sub-projects 45 | 46 | ** Quantitative Algebraic Theories 47 | 48 | *** Iteration Theories 49 | https://personal.cis.strath.ac.uk/r.mardare/papers/IterationTh.pdf 50 | 51 | ** Constructive partiality combined with non-determinism and probability 52 | 53 | Objective: to make a partial probability monad that works 54 | constructively. This is so we can model languages with recursion 55 | (so partiality) and probabilistic behaviour, such as 56 | 57 | Plan: to combine the "free omegaCPO" monad (defined using 58 | inductive-inductive definitions) with the probability monad 59 | generated from the quantitaive algebraic theory. This will live in 60 | the category of metric spaces and non-expansive maps. 61 | 62 | Related work: 63 | - *Combining Nondeterminism, Probability, and Termination: 64 | Equational and Metric Reasoning* by Matteo Mio, Ralph Sarkis, 65 | Valeria Vignudelli. https://arxiv.org/pdf/2012.00382.pdf 66 | 67 | This carries out the above programme in a classical setting, so 68 | the partiality is modelled simply by `_+1`. Includes the correct 69 | categorical definitions though. 70 | 71 | ** Searching spaces for global maxima/minima 72 | 73 | Objective: searching for global maxima/minima for functions 74 | 75 | Plan: to reuse the integration monad of O'Connor and Spitters to 76 | allowing systematic searching for global minima and maxima, which 77 | will allow for optmisation problems to be solved exactly (up to 78 | some level of approximation). 79 | 80 | Related work: 81 | - LICS'21 paper by Todd Ambridge Waugh and Dan Ghica. 82 | -------------------------------------------------------------------------------- /shell.nix: -------------------------------------------------------------------------------- 1 | { pkgs ? import { } }: 2 | with pkgs; 3 | mkShell { 4 | buildInputs = [ 5 | (agda.withPackages (ps: [ 6 | ps.standard-library 7 | ])) 8 | ]; 9 | } 10 | -------------------------------------------------------------------------------- /src/Algebra/MonoidOfSemigroup.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level using (_⊔_) 4 | open import Algebra 5 | 6 | module Algebra.MonoidOfSemigroup {m₁ m₂} (M : CommutativeSemigroup m₁ m₂) where 7 | 8 | open import Data.Product using (_,_) 9 | open import Relation.Binary 10 | 11 | open CommutativeSemigroup M 12 | 13 | data MCarrier : Set m₁ where 14 | u : MCarrier 15 | `_ : Carrier → MCarrier 16 | 17 | _⊕_ : MCarrier → MCarrier → MCarrier 18 | u ⊕ y = y 19 | (` x) ⊕ u = ` x 20 | (` x) ⊕ (` y) = ` (x ∙ y) 21 | 22 | data _≃_ : MCarrier → MCarrier → Set (m₁ ⊔ m₂) where 23 | u≃u : u ≃ u 24 | x≃y : ∀ {x y} → x ≈ y → (` x) ≃ (` y) 25 | 26 | ≃-refl : ∀ {x} → x ≃ x 27 | ≃-refl {u} = u≃u 28 | ≃-refl {` x} = x≃y refl 29 | 30 | ≃-sym : ∀ {x y} → x ≃ y → y ≃ x 31 | ≃-sym u≃u = u≃u 32 | ≃-sym (x≃y e) = x≃y (sym e) 33 | 34 | ≃-trans : ∀ {x y z} → x ≃ y → y ≃ z → x ≃ z 35 | ≃-trans u≃u u≃u = u≃u 36 | ≃-trans (x≃y eq₁) (x≃y eq₂) = x≃y (trans eq₁ eq₂) 37 | 38 | ≃-isEquivalence : IsEquivalence _≃_ 39 | ≃-isEquivalence = record { refl = ≃-refl ; sym = ≃-sym ; trans = ≃-trans } 40 | 41 | ≃-setoid : Setoid m₁ (m₁ ⊔ m₂) 42 | ≃-setoid = record { isEquivalence = ≃-isEquivalence } 43 | 44 | embedding : ∀ {x y} → (` x) ≃ (` y) → x ≈ y 45 | embedding (x≃y e) = e 46 | 47 | ⊕-cong : Congruent₂ _≃_ _⊕_ 48 | ⊕-cong u≃u y₁≃y₂ = y₁≃y₂ 49 | ⊕-cong (x≃y x₁≃x₂) u≃u = x≃y x₁≃x₂ 50 | ⊕-cong (x≃y x₁≃x₂) (x≃y y₁≃y₂) = x≃y (∙-cong x₁≃x₂ y₁≃y₂) 51 | 52 | ⊕-assoc : Associative _≃_ _⊕_ 53 | ⊕-assoc u _ _ = ≃-refl 54 | ⊕-assoc (` _) u _ = ≃-refl 55 | ⊕-assoc (` _) (` _) u = ≃-refl 56 | ⊕-assoc (` x) (` y) (` z) = x≃y (assoc x y z) 57 | 58 | ⊕-comm : Commutative _≃_ _⊕_ 59 | ⊕-comm u u = u≃u 60 | ⊕-comm u (` x) = ≃-refl 61 | ⊕-comm (` x) u = ≃-refl 62 | ⊕-comm (` x) (` y) = x≃y (comm x y) 63 | 64 | +-identityʳ : RightIdentity _≃_ u _⊕_ 65 | +-identityʳ u = ≃-refl 66 | +-identityʳ (` x) = ≃-refl 67 | 68 | +-identityˡ : LeftIdentity _≃_ u _⊕_ 69 | +-identityˡ u = ≃-refl 70 | +-identityˡ (` x) = ≃-refl 71 | 72 | +-identity : Identity _≃_ u _⊕_ 73 | +-identity = +-identityˡ , +-identityʳ 74 | 75 | ------------------------------------------------------------------------ 76 | -- Algebraic Structures 77 | 78 | ⊕-isMagma : IsMagma _≃_ _⊕_ 79 | ⊕-isMagma = record 80 | { isEquivalence = ≃-isEquivalence 81 | ; ∙-cong = ⊕-cong 82 | } 83 | 84 | ⊕-isSemigroup : IsSemigroup _≃_ _⊕_ 85 | ⊕-isSemigroup = record 86 | { isMagma = ⊕-isMagma 87 | ; assoc = ⊕-assoc 88 | } 89 | 90 | ⊕-u-isMonoid : IsMonoid _≃_ _⊕_ u 91 | ⊕-u-isMonoid = record 92 | { isSemigroup = ⊕-isSemigroup 93 | ; identity = +-identity 94 | } 95 | 96 | ⊕-u-isCommutativeMonoid : IsCommutativeMonoid _≃_ _⊕_ u 97 | ⊕-u-isCommutativeMonoid = record 98 | { isMonoid = ⊕-u-isMonoid 99 | ; comm = ⊕-comm 100 | } 101 | 102 | ------------------------------------------------------------------------ 103 | -- Algebraic bundles 104 | 105 | ⊕-magma : Magma m₁ (m₁ ⊔ m₂) 106 | ⊕-magma = record 107 | { isMagma = ⊕-isMagma 108 | } 109 | 110 | ⊕-semigroup : Semigroup m₁ (m₁ ⊔ m₂) 111 | ⊕-semigroup = record 112 | { isSemigroup = ⊕-isSemigroup 113 | } 114 | 115 | ⊕-u-monoid : Monoid m₁ (m₁ ⊔ m₂) 116 | ⊕-u-monoid = record 117 | { isMonoid = ⊕-u-isMonoid 118 | } 119 | 120 | ⊕-u-commutativeMonoid : CommutativeMonoid m₁ (m₁ ⊔ m₂) 121 | ⊕-u-commutativeMonoid = record 122 | { isCommutativeMonoid = ⊕-u-isCommutativeMonoid 123 | } 124 | -------------------------------------------------------------------------------- /src/Algebra/Solver/CommutativeSemigroup.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra 4 | 5 | module Algebra.Solver.CommutativeSemigroup {m₁ m₂} (G : CommutativeSemigroup m₁ m₂) where 6 | 7 | open import Data.Maybe as Maybe 8 | import Relation.Binary.PropositionalEquality as PropositionalEquality 9 | 10 | open import Data.Nat using (ℕ; suc) 11 | open import Data.Fin using (Fin; #_) 12 | open import Data.Vec as Vec using (Vec) 13 | open import Data.Vec using (_∷_; []) public 14 | open import Data.Vec.Properties using (lookup-map) 15 | open import Relation.Nullary.Decidable.Core using (True) 16 | 17 | import Algebra.MonoidOfSemigroup (G) as M 18 | import Algebra.Solver.CommutativeMonoid 19 | module Solver = Algebra.Solver.CommutativeMonoid (M.⊕-u-commutativeMonoid) 20 | 21 | open CommutativeSemigroup G using () renaming (_≈_ to _≈'_; Carrier to InnerCarrier; _∙_ to _∙'_) 22 | open CommutativeMonoid M.⊕-u-commutativeMonoid using (Carrier; _≈_; reflexive; setoid; _∙_; ∙-cong; refl; sym) 23 | open M using (`_; embedding) public 24 | open Solver using (Expr; ⟦_⟧; normalise; module R; ⟦_⟧⇓; _≟_) 25 | 26 | lift-env : ∀ {n} → Vec InnerCarrier n → Vec Carrier n 27 | lift-env = Vec.map (`_) 28 | 29 | data Expr' : ℕ → Set where 30 | var : ∀ {n} → Fin n → Expr' n 31 | _⊕_ : ∀ {n} → Expr' n → Expr' n → Expr' n 32 | 33 | embed-expr : ∀ {n} → Expr' n → Expr n 34 | embed-expr (var x) = Solver.var x 35 | embed-expr (e₁ ⊕ e₂) = embed-expr e₁ Solver.⊕ embed-expr e₂ 36 | 37 | ⟦_⟧' : ∀ {n} → Expr' n → Vec InnerCarrier n → InnerCarrier 38 | ⟦ var x ⟧' ρ = Vec.lookup ρ x 39 | ⟦ e₁ ⊕ e₂ ⟧' ρ = ⟦ e₁ ⟧' ρ ∙' ⟦ e₂ ⟧' ρ 40 | 41 | lem : ∀ {n} (e : Expr' n) (ρ : Vec InnerCarrier n) → ⟦ embed-expr e ⟧ (lift-env ρ) ≈ ` (⟦ e ⟧' ρ) 42 | lem (var x) ρ = reflexive (lookup-map x `_ ρ) 43 | lem (e ⊕ e₁) ρ = begin 44 | ⟦ embed-expr e Solver.⊕ embed-expr e₁ ⟧ (lift-env ρ) 45 | ≈⟨ refl ⟩ 46 | ⟦ embed-expr e ⟧ (lift-env ρ) ∙ ⟦ embed-expr e₁ ⟧ (lift-env ρ) 47 | ≈⟨ ∙-cong (lem e ρ) (lem e₁ ρ) ⟩ 48 | ` (⟦ e ⟧' ρ ∙' ⟦ e₁ ⟧' ρ) 49 | ∎ 50 | where open import Relation.Binary.Reasoning.Setoid (setoid) 51 | 52 | injective : ∀ {n}(e₁ e₂ : Expr' n)(ρ : Vec (G .CommutativeSemigroup.Carrier) n) → 53 | ⟦ embed-expr e₁ ⟧ (lift-env ρ) ≈ ⟦ embed-expr e₂ ⟧ (lift-env ρ) → 54 | ⟦ e₁ ⟧' ρ ≈' ⟦ e₂ ⟧' ρ 55 | injective e₁ e₂ ρ embed-e₁≈embed-e₂ = 56 | embedding (begin 57 | ` ⟦ e₁ ⟧' ρ ≈⟨ sym (lem e₁ ρ) ⟩ 58 | ⟦ embed-expr e₁ ⟧ (lift-env ρ) ≈⟨ embed-e₁≈embed-e₂ ⟩ 59 | ⟦ embed-expr e₂ ⟧ (lift-env ρ) ≈⟨ lem e₂ ρ ⟩ 60 | ` ⟦ e₂ ⟧' ρ 61 | ∎) 62 | where open import Relation.Binary.Reasoning.Setoid (setoid) 63 | 64 | prove′ : ∀ {n} (e₁ e₂ : Expr' n) → Maybe (∀ ρ → ⟦ e₁ ⟧' ρ ≈' ⟦ e₂ ⟧' ρ) 65 | prove′ e₁ e₂ = 66 | Maybe.map lemma (decToMaybe (normalise (embed-expr e₁) ≟ normalise (embed-expr e₂))) 67 | where 68 | open PropositionalEquality using (_≡_; cong) 69 | open PropositionalEquality.≡-Reasoning 70 | lemma : normalise (embed-expr e₁) ≡ normalise (embed-expr e₂) → ∀ ρ → ⟦ e₁ ⟧' ρ ≈' ⟦ e₂ ⟧' ρ 71 | lemma eq ρ = 72 | injective e₁ e₂ ρ 73 | (R.prove (lift-env ρ) (embed-expr e₁) (embed-expr e₂) (reflexive (begin 74 | ⟦ normalise (embed-expr e₁) ⟧⇓ (lift-env ρ) ≡⟨ cong (λ e → ⟦ e ⟧⇓ (lift-env ρ)) eq ⟩ 75 | ⟦ normalise (embed-expr e₂) ⟧⇓ (lift-env ρ) ∎))) 76 | 77 | prove : ∀ n (e₁ e₂ : Expr' n) → From-just (prove′ e₁ e₂) 78 | prove _ e₁ e₂ = from-just (prove′ e₁ e₂) 79 | 80 | import Data.Nat.Properties as ℕₚ 81 | 82 | v# : ∀ m {n} {m_ -} _≰_ {- _≱_ _≮_ _≯_ -} 35 | infix 8 1/_ _/2 36 | infixl 7 _*_ 37 | infixl 6 _+_ 38 | 39 | fog : ℚ⁺ → ℚ 40 | fog q = q .rational 41 | 42 | gof : ∀ q → q ℚ.> 0ℚ → ℚ⁺ 43 | gof q q>0 .rational = q 44 | gof q q>0 .positive = ℚ.positive q>0 45 | 46 | _+_ : ℚ⁺ → ℚ⁺ → ℚ⁺ 47 | (q + r) .rational = q .rational ℚ.+ r .rational 48 | (q + r) .positive = ℚ.positive (ℚ.+-mono-< (ℚ.positive⁻¹ (q .positive)) (ℚ.positive⁻¹ (r .positive))) 49 | 50 | _*_ : ℚ⁺ → ℚ⁺ → ℚ⁺ 51 | (q * r) .rational = q .rational ℚ.* r .rational 52 | (q * r) .positive = 53 | ℚ.positive (ℚ.≤-<-trans (ℚ.≤-reflexive (ℚ.≃-sym (ℚ.*-zeroʳ (q .rational)))) 54 | (ℚ.*-monoʳ-<-pos {q .rational} (q .positive) (ℚ.positive⁻¹ {r .rational} (r .positive)))) 55 | 56 | positive⇒nonzero : ∀ {p} → ℚ.Positive p → ℤ.∣ ↥ p ∣ ℚ.≢0 57 | positive⇒nonzero {ℚ.mkℚᵘ (+_ (suc n)) denominator-1} +ve = tt 58 | 59 | 1/-positive : ∀ p → (+ve : ℚ.Positive p) → ℚ.Positive (ℚ.1/_ p {positive⇒nonzero {p} +ve}) 60 | 1/-positive (ℚ.mkℚᵘ (+_ (suc n)) d) tt = tt 61 | 62 | 1/_ : ℚ⁺ → ℚ⁺ 63 | (1/ q) .rational = ℚ.1/_ (q .rational) {positive⇒nonzero {q .rational} (q .positive)} 64 | (1/ q) .positive = 1/-positive (q .rational) (q .positive) 65 | 66 | ½ : ℚ⁺ 67 | ½ .rational = ℚ.½ 68 | ½ .positive = tt 69 | 70 | 1ℚ⁺ : ℚ⁺ 71 | 1ℚ⁺ .rational = 1ℚ 72 | 1ℚ⁺ .positive = tt 73 | 74 | _/2 : ℚ⁺ → ℚ⁺ 75 | q /2 = ½ * q 76 | 77 | data _≃_ : ℚ⁺ → ℚ⁺ → Set where 78 | r≃r : ∀ {q r} → q .rational ℚ.≃ r .rational → q ≃ r 79 | 80 | data _≤_ : ℚ⁺ → ℚ⁺ → Set where 81 | r≤r : ∀ {q r} → q .rational ℚ.≤ r .rational → q ≤ r 82 | 83 | data _<_ : ℚ⁺ → ℚ⁺ → Set where 84 | r (λ x₁ → q≰r (r≤r x₁)))) 166 | 167 | ------------------------------------------------------------------------ 168 | -- Deciding the order and ⊓ and ⊔ 169 | open import Relation.Nullary using (yes; no; Dec) 170 | import Relation.Nullary.Decidable as Dec 171 | open import Data.Bool.Base using (Bool; if_then_else_) 172 | 173 | _≤?_ : (q r : ℚ⁺) → Dec (q ≤ r) 174 | q ≤? r with q .rational ℚ.≤? r .rational 175 | ... | yes q≤r = yes (r≤r q≤r) 176 | ... | no q≰r = no (λ { (r≤r q≤r) → q≰r q≤r }) 177 | 178 | _⊓_ : ℚ⁺ → ℚ⁺ → ℚ⁺ 179 | q ⊓ r with q ≤? r 180 | ... | yes _ = q 181 | ... | no _ = r 182 | 183 | ⊓-lower-1 : ∀ q r → (q ⊓ r) ≤ q 184 | ⊓-lower-1 q r with q ≤? r 185 | ... | yes _ = ≤-refl 186 | ... | no q≰r = ≰⇒≥ q≰r 187 | 188 | ⊓-lower-2 : ∀ q r → (q ⊓ r) ≤ r 189 | ⊓-lower-2 q r with q ≤? r 190 | ... | yes q≤r = q≤r 191 | ... | no q≰r = ≤-refl 192 | 193 | ⊓-selective : ∀ q r → (q ⊓ r) ≃ q ⊎ (q ⊓ r) ≃ r 194 | ⊓-selective q r with q ≤? r 195 | ... | yes _ = inj₁ ≃-refl 196 | ... | no _ = inj₂ ≃-refl 197 | 198 | _⊔_ : ℚ⁺ → ℚ⁺ → ℚ⁺ 199 | q ⊔ r with q ≤? r 200 | ... | yes _ = r 201 | ... | no _ = q 202 | 203 | ⊔-upper-1 : ∀ q r → q ≤ (q ⊔ r) 204 | ⊔-upper-1 q r with q ≤? r 205 | ... | yes q≤r = q≤r 206 | ... | no q≰r = ≤-refl 207 | 208 | ⊔-upper-2 : ∀ q r → r ≤ (q ⊔ r) 209 | ⊔-upper-2 q r with q ≤? r 210 | ... | yes q≤r = ≤-refl 211 | ... | no q≰r = ≰⇒≥ q≰r 212 | 213 | ------------------------------------------------------------------------ 214 | -- Properties of _+_ 215 | 216 | +-cong : Congruent₂ _≃_ _+_ 217 | +-cong (r≃r q₁≃q₂) (r≃r r₁≃r₂) = r≃r (ℚ.+-cong q₁≃q₂ r₁≃r₂) 218 | 219 | +-congʳ : ∀ p {q r} → q ≃ r → p + q ≃ p + r 220 | +-congʳ p q≃r = +-cong (≃-refl {p}) q≃r 221 | 222 | +-congˡ : ∀ p {q r} → q ≃ r → q + p ≃ r + p 223 | +-congˡ p q≃r = +-cong q≃r (≃-refl {p}) 224 | 225 | +-comm : Commutative _≃_ _+_ 226 | +-comm q r = r≃r (ℚ.+-comm (q .rational) (r .rational)) 227 | 228 | +-assoc : Associative _≃_ _+_ 229 | +-assoc q r s = r≃r (ℚ.+-assoc (q .rational) (r .rational) (s .rational)) 230 | 231 | +-mono-≤ : _+_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_ 232 | +-mono-≤ (r≤r x≤y) (r≤r u≤v) = r≤r (ℚ.+-mono-≤ x≤y u≤v) 233 | 234 | +-increasing : ∀ {q r} → q ≤ q + r 235 | +-increasing {q}{r} = r≤r (ℚ.p≤p+q {q .rational}{r .rational} (ℚ.positive⇒nonNegative {r .rational} (r .positive))) 236 | 237 | -- private 238 | -- blah : ∀ {q r} → q ℚ.≠ r → q ℚ.≤ r → q ℚ.< r 239 | -- blah q≠r (ℚ.*≤* q≤r) = ℚ.*<* (ℤ.≤∧≢⇒< q≤r (λ x → q≠r (ℚ.*≡* x) )) 240 | 241 | {- 242 | postulate -- FIXME 243 | ≤-split : ∀ {ε₁ ε₂} → ε₁ ≤ ε₂ → (ε₁ ≃ ε₂) ⊎ (Σ[ δ ∈ ℚ⁺ ] (ε₁ + δ ≃ ε₂)) 244 | -- ≤-split {⟨ ε₁ , _ ⟩}{⟨ ε₂ , _ ⟩} (r≤r ε₁≤ε₂) with ε₁ ℚᵘ.≃? ε₂ 245 | -- ... | yes ε₁≃ε₂ = inj₁ (r≃r ε₁≃ε₂) -- ⟨ ε₂ .rational ℚᵘ.- ε₁ .rational , ℚᵘ.positive {!ℚ.p≤q⇒0≤q-p!} ⟩ , {!!} 246 | -- ... | no ε₁≠ε₂ = inj₂ (⟨ ε₂ ℚᵘ.- ε₁ , ℚᵘ.positive (blah {!!} (ℚ.p≤q⇒0≤q-p ε₁≤ε₂)) ⟩ , r≃r {!!}) 247 | -} 248 | ------------------------------------------------------------------------ 249 | -- Properties of _*_ 250 | 251 | *-cong : Congruent₂ _≃_ _*_ 252 | *-cong (r≃r q₁≃q₂) (r≃r r₁≃r₂) = r≃r (ℚ.*-cong q₁≃q₂ r₁≃r₂) 253 | 254 | *-congʳ : ∀ p {q r} → q ≃ r → p * q ≃ p * r 255 | *-congʳ p q≃r = *-cong (≃-refl {p}) q≃r 256 | 257 | *-congˡ : ∀ p {q r} → q ≃ r → q * p ≃ r * p 258 | *-congˡ p q≃r = *-cong q≃r (≃-refl {p}) 259 | 260 | *-mono-≤ : _*_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_ 261 | *-mono-≤ {x}{y}{u}{v} (r≤r x≤y) (r≤r u≤v) = 262 | r≤r (ℚ.≤-trans (ℚ.*-monoʳ-≤-pos {x .rational} (x .positive) u≤v) 263 | (ℚ.*-monoˡ-≤-pos (v .positive) x≤y)) 264 | 265 | *-mono-< : _*_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_ 266 | *-mono-< {x}{y}{u}{v} (r 0ℚ → a ℚ.≤ b ℚ.+ ε) → a ℚ.≤ b 218 | closed₁ a b h with a ℚ.≤? b 219 | ... | yes a≤b = a≤b 220 | ... | no ¬a≤b with ℚ.≰⇒> ¬a≤b 221 | ... | b0 → begin 248 | a ≤⟨ h (gof ε ε>0) ⟩ 249 | fog (b ℚ⁺.+ gof ε ε>0) ≈⟨ ℚ.*≡* refl ⟩ 250 | fog b ℚ.+ ε ∎ 251 | where open ℚ.≤-Reasoning 252 | 253 | rational : ℚ → ℝᵘ -- FIXME: should be ℚ≥0 254 | rational r .contains q = r ℚ.≤ ℚ⁺.fog q 255 | rational r .upper q₁≤q₂ r≤q₁ = ℚ.≤-trans r≤q₁ (ℚ⁺.fog-mono q₁≤q₂) 256 | rational r .closed {q} = closedness.closed' r q 257 | 258 | {- 259 | -- essentially this holds because closure is a monad, and we are 260 | -- taking the free algebra here. 261 | rational-alt : ℚ → ℝᵘ 262 | rational-alt r .contains q = ∀ ε → r ℚ.≤ ℚ⁺.fog (q ℚ⁺.+ ε) 263 | rational-alt r .upper {q₁}{q₂} q₁≤q₂ r∈q₁ ε = 264 | begin 265 | r 266 | ≤⟨ r∈q₁ ε ⟩ 267 | ℚ⁺.fog (q₁ ℚ⁺.+ ε) 268 | ≤⟨ ℚ⁺.fog-mono (ℚ⁺.+-mono-≤ q₁≤q₂ (ℚ⁺.≤-refl {ε})) ⟩ 269 | ℚ⁺.fog (q₂ ℚ⁺.+ ε) 270 | ∎ 271 | where open ℚ.≤-Reasoning 272 | rational-alt r .closed {ε} = {!!} 273 | -} 274 | 275 | ------------------------------------------------------------------------------ 276 | -- Generic construction and proofs for binary operators 277 | module binary-op (_⚈_ : ℚ⁺ → ℚ⁺ → ℚ⁺) (⚈-comm : Commutative ℚ⁺._≃_ _⚈_) where 278 | 279 | _⚈ℝ_ : ℝᵘ → ℝᵘ → ℝᵘ 280 | (x ⚈ℝ y) .contains q = ∀ s → Σ[ q₁ ∈ ℚ⁺ ] Σ[ q₂ ∈ ℚ⁺ ] (q₁ ⚈ q₂ ℚ⁺.≤ q ℚ⁺.+ s × x .contains q₁ × y .contains q₂) 281 | (x ⚈ℝ y) .upper q₁≤q₂ x⚈y s = 282 | let q'₁ , q'₂ , ineq , x-q'₁ , y-q'₂ = x⚈y s in 283 | q'₁ , q'₂ , ℚ⁺.≤-trans ineq (ℚ⁺.+-mono-≤ q₁≤q₂ ℚ⁺.≤-refl) , x-q'₁ , y-q'₂ 284 | (x ⚈ℝ y) .closed {q} h s = 285 | let q₁ , q₂ , ineq , x-q₁ , y-q₂ = h (s ℚ⁺./2) (s ℚ⁺./2) in 286 | q₁ , q₂ , ℚ⁺.≤-trans ineq (ℚ⁺.≤-reflexive (ℚ⁺.≃-trans (ℚ⁺.+-assoc q (s ℚ⁺./2) (s ℚ⁺./2)) (ℚ⁺.+-congʳ q ℚ⁺.half+half))) , x-q₁ , y-q₂ 287 | 288 | mono-≤ : _⚈ℝ_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_ 289 | mono-≤ x≤y u≤v .*≤* y⚈v s = 290 | let q₁ , q₂ , ineq , y-q₁ , v-q₂ = y⚈v s in 291 | q₁ , q₂ , ineq , x≤y .*≤* y-q₁ , u≤v .*≤* v-q₂ 292 | 293 | cong : Congruent₂ _≃_ _⚈ℝ_ 294 | cong {x}{y}{u}{v} x≃y u≃v = 295 | mono-≤ {x}{y}{u}{v} (x≃y .proj₁) (u≃v .proj₁) , 296 | mono-≤ {y}{x}{v}{u} (x≃y .proj₂) (u≃v .proj₂) 297 | 298 | comm : Commutative _≃_ _⚈ℝ_ 299 | comm x y .proj₁ .*≤* x⚈y s = 300 | let q₁ , q₂ , ineq , x-q₁ , y-q₂ = x⚈y s in 301 | q₂ , q₁ , ℚ⁺.≤-trans (ℚ⁺.≤-reflexive (⚈-comm q₂ q₁)) ineq , y-q₂ , x-q₁ 302 | comm x y .proj₂ .*≤* x⚈y s = 303 | let q₁ , q₂ , ineq , x-q₁ , y-q₂ = x⚈y s in 304 | q₂ , q₁ , ℚ⁺.≤-trans (ℚ⁺.≤-reflexive (⚈-comm q₂ q₁)) ineq , y-q₂ , x-q₁ 305 | 306 | ------------------------------------------------------------------------------ 307 | -- Addition 308 | open binary-op (ℚ⁺._+_) (ℚ⁺.+-comm) 309 | renaming ( _⚈ℝ_ to _+_ 310 | ; comm to +-comm 311 | ; mono-≤ to +-mono-≤ 312 | ; cong to +-cong ) 313 | public 314 | 315 | +-identityˡ : LeftIdentity _≃_ 0ℝ _+_ 316 | +-identityˡ x .proj₁ .*≤* {q} x-q s = 317 | s , q , ℚ⁺.≤-reflexive (ℚ⁺.+-comm s q) , tt , x-q 318 | +-identityˡ x .proj₂ .*≤* {q} 0+x = 319 | x .closed (λ s → let q₁ , q₂ , ineq , tt , x-q₂ = 0+x s in 320 | x .upper (ℚ⁺.≤-trans (ℚ⁺.≤-trans ℚ⁺.+-increasing (ℚ⁺.≤-reflexive (ℚ⁺.+-comm q₂ q₁))) ineq) x-q₂) 321 | 322 | +-identityʳ : RightIdentity _≃_ 0ℝ _+_ 323 | +-identityʳ x = ≃-trans (+-comm x 0ℝ) (+-identityˡ x) 324 | 325 | +-identity : Identity _≃_ 0ℝ _+_ 326 | +-identity = +-identityˡ , +-identityʳ 327 | 328 | +-assoc : Associative _≃_ _+_ 329 | +-assoc x y z .proj₁ .*≤* {q} x+⟨y+z⟩ s = 330 | let q₁ , q₂ , q₁+q₂≤q+s/2 , x-q₁ , y+z = x+⟨y+z⟩ (s ℚ⁺./2) in 331 | let q₃ , q₄ , q₃+q₄≤q₂+s/2 , y-q₃ , z-q₄ = y+z (s ℚ⁺./2) in 332 | q₁ ℚ⁺.+ q₃ , q₄ , 333 | (begin 334 | (q₁ ℚ⁺.+ q₃) ℚ⁺.+ q₄ 335 | ≈⟨ ℚ⁺.+-assoc q₁ q₃ q₄ ⟩ 336 | q₁ ℚ⁺.+ (q₃ ℚ⁺.+ q₄) 337 | ≤⟨ ℚ⁺.+-mono-≤ ℚ⁺.≤-refl q₃+q₄≤q₂+s/2 ⟩ 338 | q₁ ℚ⁺.+ (q₂ ℚ⁺.+ s ℚ⁺./2) 339 | ≈⟨ ℚ⁺.≃-sym (ℚ⁺.+-assoc q₁ q₂ (s ℚ⁺./2)) ⟩ 340 | (q₁ ℚ⁺.+ q₂) ℚ⁺.+ s ℚ⁺./2 341 | ≤⟨ ℚ⁺.+-mono-≤ q₁+q₂≤q+s/2 ℚ⁺.≤-refl ⟩ 342 | (q ℚ⁺.+ s ℚ⁺./2) ℚ⁺.+ s ℚ⁺./2 343 | ≈⟨ ℚ⁺.+-assoc q (s ℚ⁺./2) (s ℚ⁺./2) ⟩ 344 | q ℚ⁺.+ (s ℚ⁺./2 ℚ⁺.+ s ℚ⁺./2) 345 | ≈⟨ ℚ⁺.+-congʳ q ℚ⁺.half+half ⟩ 346 | q ℚ⁺.+ s 347 | ∎) , 348 | (λ s₁ → q₁ , q₃ , ℚ⁺.+-increasing , x-q₁ , y-q₃) , z-q₄ 349 | where open ℚ⁺.≤-Reasoning 350 | +-assoc x y z .proj₂ .*≤* {q} ⟨x+y⟩+z s = 351 | let q₁ , q₂ , q₁+q₂≤q+s/2 , x+y , z-q₂ = ⟨x+y⟩+z (s ℚ⁺./2) in 352 | let q₃ , q₄ , q₃+q₄≤q₁+s/2 , x-q₃ , y-q₄ = x+y (s ℚ⁺./2) in 353 | q₃ , q₄ ℚ⁺.+ q₂ , 354 | (begin 355 | q₃ ℚ⁺.+ (q₄ ℚ⁺.+ q₂) 356 | ≈⟨ ℚ⁺.≃-sym (ℚ⁺.+-assoc q₃ q₄ q₂) ⟩ 357 | (q₃ ℚ⁺.+ q₄) ℚ⁺.+ q₂ 358 | ≤⟨ ℚ⁺.+-mono-≤ q₃+q₄≤q₁+s/2 ℚ⁺.≤-refl ⟩ 359 | (q₁ ℚ⁺.+ s ℚ⁺./2) ℚ⁺.+ q₂ 360 | ≈⟨ ℚ⁺.+-assoc q₁ (s ℚ⁺./2) q₂ ⟩ 361 | q₁ ℚ⁺.+ (s ℚ⁺./2 ℚ⁺.+ q₂) 362 | ≈⟨ ℚ⁺.+-congʳ q₁ (ℚ⁺.+-comm (s ℚ⁺./2) q₂) ⟩ 363 | q₁ ℚ⁺.+ (q₂ ℚ⁺.+ s ℚ⁺./2) 364 | ≈⟨ ℚ⁺.≃-sym (ℚ⁺.+-assoc q₁ q₂ (s ℚ⁺./2)) ⟩ 365 | (q₁ ℚ⁺.+ q₂) ℚ⁺.+ s ℚ⁺./2 366 | ≤⟨ ℚ⁺.+-mono-≤ q₁+q₂≤q+s/2 ℚ⁺.≤-refl ⟩ 367 | (q ℚ⁺.+ s ℚ⁺./2) ℚ⁺.+ s ℚ⁺./2 368 | ≈⟨ ℚ⁺.+-assoc q (s ℚ⁺./2) (s ℚ⁺./2) ⟩ 369 | q ℚ⁺.+ (s ℚ⁺./2 ℚ⁺.+ s ℚ⁺./2) 370 | ≈⟨ ℚ⁺.+-congʳ q ℚ⁺.half+half ⟩ 371 | q ℚ⁺.+ s 372 | ∎) , 373 | x-q₃ , 374 | λ s₁ → q₄ , q₂ , ℚ⁺.+-increasing , y-q₄ , z-q₂ 375 | where open ℚ⁺.≤-Reasoning 376 | 377 | +-increasingʳ : ∀ x y → x ≤ x + y 378 | +-increasingʳ x y = 379 | begin 380 | x ≈⟨ ≃-sym (+-identityʳ x) ⟩ 381 | x + 0ℝ ≤⟨ +-mono-≤ ≤-refl (0-least y) ⟩ 382 | x + y ∎ 383 | where open ≤-Reasoning 384 | 385 | +-increasingˡ : ∀ x y → y ≤ x + y 386 | +-increasingˡ x y = 387 | begin 388 | y ≈⟨ ≃-sym (+-identityˡ y) ⟩ 389 | 0ℝ + y ≤⟨ +-mono-≤ (0-least x) ≤-refl ⟩ 390 | x + y ∎ 391 | where open ≤-Reasoning 392 | 393 | +-infinityʳ : ∀ x → (x + ∞) ≃ ∞ 394 | +-infinityʳ x .proj₁ = ∞-greatest _ 395 | +-infinityʳ x .proj₂ .*≤* {ε} x+∞∋ε = 396 | let _ , _ , _ , _ , ϕ = x+∞∋ε ℚ⁺.1ℚ⁺ in 397 | ϕ 398 | 399 | ------------------------------------------------------------------------------ 400 | -- Structures for the algebraic structure of _+_ 401 | +-isMagma : IsMagma _≃_ _+_ 402 | +-isMagma = record { isEquivalence = ≃-isEquivalence ; ∙-cong = +-cong } 403 | 404 | +-isSemigroup : IsSemigroup _≃_ _+_ 405 | +-isSemigroup = record { isMagma = +-isMagma ; assoc = +-assoc } 406 | 407 | +-isCommutativeSemigroup : IsCommutativeSemigroup _≃_ _+_ 408 | +-isCommutativeSemigroup = record { isSemigroup = +-isSemigroup ; comm = +-comm } 409 | 410 | +-0-isMonoid : IsMonoid _≃_ _+_ 0ℝ 411 | +-0-isMonoid = record { isSemigroup = +-isSemigroup ; identity = +-identity } 412 | 413 | +-0-isCommutativeMonoid : IsCommutativeMonoid _≃_ _+_ 0ℝ 414 | +-0-isCommutativeMonoid = record { isMonoid = +-0-isMonoid ; comm = +-comm } 415 | 416 | ------------------------------------------------------------------------------ 417 | -- Bundles 418 | +-commutativeSemigroup : CommutativeSemigroup (suc 0ℓ) 0ℓ 419 | +-commutativeSemigroup = record { isCommutativeSemigroup = +-isCommutativeSemigroup } 420 | 421 | open interchange +-commutativeSemigroup public 422 | 423 | ------------------------------------------------------------------------------ 424 | -- Multiplication 425 | open binary-op (ℚ⁺._*_) (ℚ⁺.*-comm) 426 | renaming ( _⚈ℝ_ to _*_ 427 | ; comm to *-comm 428 | ; mono-≤ to *-mono-≤ 429 | ; cong to *-cong ) 430 | public 431 | 432 | 1ℝ : ℝᵘ 433 | 1ℝ = rational 1ℚ 434 | 435 | *-zeroʳ : ∀ x → x < ∞ → (x * 0ℝ) ≃ 0ℝ 436 | *-zeroʳ x (ε₁ , x∋ε₁ , _) .proj₁ .*≤* {ε} tt s = 437 | ε₁ , 438 | 1/ ε₁ ℚ⁺.* (ε ℚ⁺.+ s) , 439 | (begin 440 | ε₁ ℚ⁺.* (1/ ε₁ ℚ⁺.* (ε ℚ⁺.+ s)) 441 | ≈⟨ ℚ⁺.≃-sym (ℚ⁺.*-assoc ε₁ (1/ ε₁) (ε ℚ⁺.+ s)) ⟩ 442 | (ε₁ ℚ⁺.* 1/ ε₁) ℚ⁺.* (ε ℚ⁺.+ s) 443 | ≈⟨ ℚ⁺.*-cong (ℚ⁺.*-inverseʳ ε₁) ℚ⁺.≃-refl ⟩ 444 | ℚ⁺.1ℚ⁺ ℚ⁺.* (ε ℚ⁺.+ s) 445 | ≈⟨ ℚ⁺.*-identityˡ (ε ℚ⁺.+ s) ⟩ 446 | ε ℚ⁺.+ s 447 | ∎) , 448 | x∋ε₁ , 449 | tt 450 | where open ℚ⁺.≤-Reasoning 451 | *-zeroʳ x x<∞ .proj₂ = 0-least (x * 0ℝ) 452 | 453 | *-zeroˡ : ∀ x → x < ∞ → (0ℝ * x) ≃ 0ℝ 454 | *-zeroˡ x x<∞ = ≃-trans (*-comm 0ℝ x) (*-zeroʳ x x<∞) 455 | 456 | *-infinityʳ : ∀ x → (x * ∞) ≃ ∞ 457 | *-infinityʳ x .proj₁ = ∞-greatest _ 458 | *-infinityʳ x .proj₂ .*≤* {ε} x∞∋ε = 459 | let _ , _ , _ , _ , ϕ = x∞∋ε ℚ⁺.1ℚ⁺ in 460 | ϕ 461 | 462 | *-identityˡ : LeftIdentity _≃_ 1ℝ _*_ 463 | *-identityˡ x .proj₁ .*≤* {ε} x∋ε s = 464 | ℚ⁺.1ℚ⁺ , 465 | ε , 466 | ℚ⁺.≤-trans (ℚ⁺.≤-reflexive (ℚ⁺.*-identityˡ ε)) ℚ⁺.+-increasing , 467 | ℚ.≤-refl , 468 | x∋ε 469 | *-identityˡ x .proj₂ .*≤* {ε} 1x∋ε = 470 | x .closed λ s → 471 | let ε₁ , ε₂ , ε₁ε₂≤ε+s , 1≤ε₁ , x∋ε₂ = 1x∋ε s in 472 | x .upper 473 | (begin 474 | ε₂ 475 | ≈⟨ ℚ⁺.≃-sym (ℚ⁺.*-identityˡ ε₂) ⟩ 476 | ℚ⁺.1ℚ⁺ ℚ⁺.* ε₂ 477 | ≤⟨ ℚ⁺.*-mono-≤ (ℚ⁺.r≤r 1≤ε₁) (ℚ⁺.≤-refl {ε₂}) ⟩ 478 | ε₁ ℚ⁺.* ε₂ 479 | ≤⟨ ε₁ε₂≤ε+s ⟩ 480 | ε ℚ⁺.+ s 481 | ∎) 482 | x∋ε₂ 483 | where open ℚ⁺.≤-Reasoning 484 | 485 | *-identityʳ : RightIdentity _≃_ 1ℝ _*_ 486 | *-identityʳ x = ≃-trans (*-comm x 1ℝ) (*-identityˡ x) 487 | 488 | *-identity : Identity _≃_ 1ℝ _*_ 489 | *-identity = *-identityˡ , *-identityʳ 490 | 491 | *-assoc : Associative _≃_ _*_ 492 | *-assoc x y z .proj₁ .*≤* {ε} x[yz]∋ε s = 493 | let ε₁ , ε₂ , ε₁ε₂≤ε+s , x∋ε₁ , yz∋ε₂ = x[yz]∋ε (s /2) in 494 | let ε₂₁ , ε₂₂ , ε₂₁ε₂₂≤ε₂+s , y∋ε₂₁ , z∋ε₂₂ = yz∋ε₂ (1/ ε₁ ℚ⁺.* s /2) in 495 | ε₁ ℚ⁺.* ε₂₁ , 496 | ε₂₂ , 497 | (begin 498 | (ε₁ ℚ⁺.* ε₂₁) ℚ⁺.* ε₂₂ 499 | ≈⟨ ℚ⁺.*-assoc ε₁ ε₂₁ ε₂₂ ⟩ 500 | ε₁ ℚ⁺.* (ε₂₁ ℚ⁺.* ε₂₂) 501 | ≤⟨ ℚ⁺.*-mono-≤ (ℚ⁺.≤-refl {ε₁}) ε₂₁ε₂₂≤ε₂+s ⟩ 502 | ε₁ ℚ⁺.* (ε₂ ℚ⁺.+ (1/ ε₁ ℚ⁺.* s /2)) 503 | ≈⟨ ℚ⁺.*-distribˡ-+ ε₁ ε₂ (1/ ε₁ ℚ⁺.* s /2) ⟩ 504 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ ε₁ ℚ⁺.* (1/ ε₁ ℚ⁺.* s /2) 505 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.≃-sym (ℚ⁺.*-assoc ε₁ (1/ ε₁) (s /2))) ⟩ 506 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ (ε₁ ℚ⁺.* 1/ ε₁) ℚ⁺.* s /2 507 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-congˡ (s /2) (ℚ⁺.*-inverseʳ ε₁)) ⟩ 508 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ ℚ⁺.1ℚ⁺ ℚ⁺.* s /2 509 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-identityˡ (s /2)) ⟩ 510 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ s /2 511 | ≤⟨ ℚ⁺.+-mono-≤ ε₁ε₂≤ε+s (ℚ⁺.≤-refl {s /2}) ⟩ 512 | (ε ℚ⁺.+ s /2) ℚ⁺.+ s /2 513 | ≈⟨ ℚ⁺.+-assoc ε (s /2) (s /2) ⟩ 514 | ε ℚ⁺.+ (s /2 ℚ⁺.+ s /2) 515 | ≈⟨ ℚ⁺.+-congʳ ε ℚ⁺.half+half ⟩ 516 | ε ℚ⁺.+ s 517 | ∎) , 518 | (λ r → ε₁ , ε₂₁ , ℚ⁺.+-increasing , x∋ε₁ , y∋ε₂₁) , 519 | z∋ε₂₂ 520 | where open ℚ⁺.≤-Reasoning 521 | *-assoc x y z .proj₂ .*≤* {ε} [xy]z∋ε s = 522 | let ε₁ , ε₂ , ε₁ε₂≤ε+ , xy∋ε₁ , z∋ε₂ = [xy]z∋ε (s /2) in 523 | let ε₁₁ , ε₁₂ , ε₁₁ε₁₂≤ε₁+ , x∋ε₁₁ , y∋ε₁₂ = xy∋ε₁ (s /2 ℚ⁺.* 1/ ε₂) in 524 | ε₁₁ , 525 | ε₁₂ ℚ⁺.* ε₂ , 526 | (begin 527 | ε₁₁ ℚ⁺.* (ε₁₂ ℚ⁺.* ε₂) 528 | ≈⟨ ℚ⁺.≃-sym (ℚ⁺.*-assoc ε₁₁ ε₁₂ ε₂) ⟩ 529 | (ε₁₁ ℚ⁺.* ε₁₂) ℚ⁺.* ε₂ 530 | ≤⟨ ℚ⁺.*-mono-≤ ε₁₁ε₁₂≤ε₁+ (ℚ⁺.≤-refl {ε₂}) ⟩ 531 | (ε₁ ℚ⁺.+ (s /2 ℚ⁺.* 1/ ε₂)) ℚ⁺.* ε₂ 532 | ≈⟨ ℚ⁺.*-distribʳ-+ ε₂ ε₁ (s /2 ℚ⁺.* 1/ ε₂) ⟩ 533 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ (s /2 ℚ⁺.* 1/ ε₂) ℚ⁺.* ε₂ 534 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-assoc (s /2) (1/ ε₂) ε₂) ⟩ 535 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ s /2 ℚ⁺.* (1/ ε₂ ℚ⁺.* ε₂) 536 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-congʳ (s /2) (ℚ⁺.*-inverseˡ ε₂)) ⟩ 537 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ s /2 ℚ⁺.* ℚ⁺.1ℚ⁺ 538 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-identityʳ (s /2)) ⟩ 539 | ε₁ ℚ⁺.* ε₂ ℚ⁺.+ s /2 540 | ≤⟨ ℚ⁺.+-mono-≤ ε₁ε₂≤ε+ (ℚ⁺.≤-refl {s /2}) ⟩ 541 | (ε ℚ⁺.+ s /2) ℚ⁺.+ s /2 542 | ≈⟨ ℚ⁺.+-assoc ε (s /2) (s /2) ⟩ 543 | ε ℚ⁺.+ (s /2 ℚ⁺.+ s /2) 544 | ≈⟨ ℚ⁺.+-congʳ ε ℚ⁺.half+half ⟩ 545 | ε ℚ⁺.+ s 546 | ∎) , 547 | x∋ε₁₁ , 548 | λ r → ε₁₂ , ε₂ , ℚ⁺.+-increasing , y∋ε₁₂ , z∋ε₂ 549 | where open ℚ⁺.≤-Reasoning 550 | 551 | *-distribʳ-+ : ∀ x y z → ((y + z) * x) ≃ ((y * x) + (z * x)) 552 | *-distribʳ-+ x y z .proj₁ .*≤* {ε} yx+zx∋ε s = 553 | let ε₁ , ε₂ , ε₁+ε₂≤ε+ , yx∋ε₁ , zx∋ε₂ = yx+zx∋ε (s /2) in 554 | let ε₁₁ , ε₁₂ , ε₁₁ε₁₂≤ε₁+ , y∋ε₁₁ , x∋ε₁₂ = yx∋ε₁ (s /2 /2) in 555 | let ε₂₁ , ε₂₂ , ε₂₁ε₂₂≤ε₂+ , z∋ε₂₁ , x∋ε₂₂ = zx∋ε₂ (s /2 /2) in 556 | ε₁₁ ℚ⁺.+ ε₂₁ , 557 | ε₁₂ ℚ⁺.⊓ ε₂₂ , 558 | (begin 559 | (ε₁₁ ℚ⁺.+ ε₂₁) ℚ⁺.* (ε₁₂ ℚ⁺.⊓ ε₂₂) 560 | ≈⟨ ℚ⁺.*-distribʳ-+ (ε₁₂ ℚ⁺.⊓ ε₂₂) ε₁₁ ε₂₁ ⟩ 561 | ε₁₁ ℚ⁺.* (ε₁₂ ℚ⁺.⊓ ε₂₂) ℚ⁺.+ ε₂₁ ℚ⁺.* (ε₁₂ ℚ⁺.⊓ ε₂₂) 562 | ≤⟨ ℚ⁺.+-mono-≤ (ℚ⁺.*-mono-≤ (ℚ⁺.≤-refl {ε₁₁}) (ℚ⁺.⊓-lower-1 ε₁₂ ε₂₂)) (ℚ⁺.*-mono-≤ (ℚ⁺.≤-refl {ε₂₁}) (ℚ⁺.⊓-lower-2 ε₁₂ ε₂₂)) ⟩ 563 | ε₁₁ ℚ⁺.* ε₁₂ ℚ⁺.+ ε₂₁ ℚ⁺.* ε₂₂ 564 | ≤⟨ ℚ⁺.+-mono-≤ ε₁₁ε₁₂≤ε₁+ ε₂₁ε₂₂≤ε₂+ ⟩ 565 | (ε₁ ℚ⁺.+ s /2 /2) ℚ⁺.+ (ε₂ ℚ⁺.+ s /2 /2) 566 | ≈⟨ ℚ⁺-interchange ε₁ (s /2 /2) ε₂ (s /2 /2) ⟩ 567 | (ε₁ ℚ⁺.+ ε₂) ℚ⁺.+ (s /2 /2 ℚ⁺.+ s /2 /2) 568 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.+ ε₂) (ℚ⁺.half+half {s /2}) ⟩ 569 | (ε₁ ℚ⁺.+ ε₂) ℚ⁺.+ s /2 570 | ≤⟨ ℚ⁺.+-mono-≤ ε₁+ε₂≤ε+ (ℚ⁺.≤-refl {s /2}) ⟩ 571 | (ε ℚ⁺.+ s /2) ℚ⁺.+ s /2 572 | ≈⟨ ℚ⁺.+-assoc ε (s /2) (s /2) ⟩ 573 | ε ℚ⁺.+ (s /2 ℚ⁺.+ s /2) 574 | ≈⟨ ℚ⁺.+-congʳ ε (ℚ⁺.half+half {s}) ⟩ 575 | ε ℚ⁺.+ s 576 | ∎) , 577 | (λ r → ε₁₁ , ε₂₁ , ℚ⁺.+-increasing , y∋ε₁₁ , z∋ε₂₁) , 578 | x∋⊓ ε₁₂ ε₂₂ x∋ε₁₂ x∋ε₂₂ 579 | where open ℚ⁺.≤-Reasoning 580 | x∋⊓ : ∀ ε₁ ε₂ → x .contains ε₁ → x .contains ε₂ → x .contains (ε₁ ℚ⁺.⊓ ε₂) 581 | x∋⊓ ε₁ ε₂ x∋ε₁ x∋ε₂ with ℚ⁺.⊓-selective ε₁ ε₂ 582 | x∋⊓ ε₁ ε₂ x∋ε₁ x∋ε₂ | inj₁ p = x .upper (ℚ⁺.≤-reflexive (ℚ⁺.≃-sym p)) x∋ε₁ 583 | x∋⊓ ε₁ ε₂ x∋ε₁ x∋ε₂ | inj₂ p = x .upper (ℚ⁺.≤-reflexive (ℚ⁺.≃-sym p)) x∋ε₂ 584 | *-distribʳ-+ x y z .proj₂ .*≤* {ε} [y+z]x∋ε s = 585 | let ε₁ , ε₂ , ε₁ε₂≤ε+ , y+z∋ε₁ , x∋ε₂ = [y+z]x∋ε (s /2) in 586 | let ε₁₁ , ε₁₂ , ε₁₁+ε₁₂≤ε₁+ , y∋ε₁₁ , z∋ε₁₂ = y+z∋ε₁ (s /2 ℚ⁺.* 1/ ε₂) in 587 | ε₁₁ ℚ⁺.* ε₂ , 588 | ε₁₂ ℚ⁺.* ε₂ , 589 | (begin 590 | (ε₁₁ ℚ⁺.* ε₂) ℚ⁺.+ (ε₁₂ ℚ⁺.* ε₂) 591 | ≈⟨ ℚ⁺.≃-sym (ℚ⁺.*-distribʳ-+ ε₂ ε₁₁ ε₁₂) ⟩ 592 | (ε₁₁ ℚ⁺.+ ε₁₂) ℚ⁺.* ε₂ 593 | ≤⟨ ℚ⁺.*-mono-≤ ε₁₁+ε₁₂≤ε₁+ (ℚ⁺.≤-refl {ε₂}) ⟩ 594 | (ε₁ ℚ⁺.+ (s /2 ℚ⁺.* 1/ ε₂)) ℚ⁺.* ε₂ 595 | ≈⟨ ℚ⁺.*-distribʳ-+ ε₂ ε₁ (s /2 ℚ⁺.* 1/ ε₂) ⟩ 596 | (ε₁ ℚ⁺.* ε₂) ℚ⁺.+ (s /2 ℚ⁺.* 1/ ε₂) ℚ⁺.* ε₂ 597 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-assoc (s /2) (1/ ε₂) ε₂) ⟩ 598 | (ε₁ ℚ⁺.* ε₂) ℚ⁺.+ s /2 ℚ⁺.* (1/ ε₂ ℚ⁺.* ε₂) 599 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-congʳ (s /2) (ℚ⁺.*-inverseˡ ε₂)) ⟩ 600 | (ε₁ ℚ⁺.* ε₂) ℚ⁺.+ s /2 ℚ⁺.* ℚ⁺.1ℚ⁺ 601 | ≈⟨ ℚ⁺.+-congʳ (ε₁ ℚ⁺.* ε₂) (ℚ⁺.*-identityʳ (s /2)) ⟩ 602 | (ε₁ ℚ⁺.* ε₂) ℚ⁺.+ s /2 603 | ≤⟨ ℚ⁺.+-mono-≤ ε₁ε₂≤ε+ (ℚ⁺.≤-refl {s /2}) ⟩ 604 | (ε ℚ⁺.+ s /2) ℚ⁺.+ s /2 605 | ≈⟨ ℚ⁺.+-assoc ε (s /2) (s /2) ⟩ 606 | ε ℚ⁺.+ (s /2 ℚ⁺.+ s /2) 607 | ≈⟨ ℚ⁺.+-congʳ ε (ℚ⁺.half+half {s}) ⟩ 608 | ε ℚ⁺.+ s 609 | ∎) , 610 | (λ r → ε₁₁ , ε₂ , ℚ⁺.+-increasing , y∋ε₁₁ , x∋ε₂) , 611 | (λ r → ε₁₂ , ε₂ , ℚ⁺.+-increasing , z∋ε₁₂ , x∋ε₂) 612 | where open ℚ⁺.≤-Reasoning 613 | 614 | *-distribˡ-+ : ∀ x y z → (x * (y + z)) ≃ ((x * y) + (x * z)) 615 | *-distribˡ-+ x y z = 616 | begin-equality 617 | x * (y + z) 618 | ≈⟨ *-comm x (y + z) ⟩ 619 | (y + z) * x 620 | ≈⟨ *-distribʳ-+ x y z ⟩ 621 | (y * x) + (z * x) 622 | ≈⟨ +-cong (*-comm y x) (*-comm z x) ⟩ 623 | (x * y) + (x * z) 624 | ∎ 625 | where open ≤-Reasoning 626 | 627 | *-distrib-+ : _DistributesOver_ _≃_ _*_ _+_ 628 | *-distrib-+ = *-distribˡ-+ , *-distribʳ-+ 629 | 630 | ------------------------------------------------------------------------------ 631 | -- Structures for the algebraic structure of _*_ 632 | 633 | *-isMagma : IsMagma _≃_ _*_ 634 | *-isMagma = record { isEquivalence = ≃-isEquivalence ; ∙-cong = *-cong } 635 | 636 | *-isSemigroup : IsSemigroup _≃_ _*_ 637 | *-isSemigroup = record { isMagma = *-isMagma ; assoc = *-assoc } 638 | 639 | *-isCommutativeSemigroup : IsCommutativeSemigroup _≃_ _*_ 640 | *-isCommutativeSemigroup = record { isSemigroup = *-isSemigroup ; comm = *-comm } 641 | 642 | *-1-isMonoid : IsMonoid _≃_ _*_ 1ℝ 643 | *-1-isMonoid = record { isSemigroup = *-isSemigroup ; identity = *-identity } 644 | 645 | *-1-isCommutativeMonoid : IsCommutativeMonoid _≃_ _*_ 1ℝ 646 | *-1-isCommutativeMonoid = record { isMonoid = *-1-isMonoid ; comm = *-comm } 647 | 648 | isSemiringWithoutAnnihilatingZero : IsSemiringWithoutAnnihilatingZero _≃_ _+_ _*_ 0ℝ 1ℝ 649 | isSemiringWithoutAnnihilatingZero = 650 | record { +-isCommutativeMonoid = +-0-isCommutativeMonoid 651 | ; *-isMonoid = *-1-isMonoid 652 | ; distrib = *-distrib-+ } 653 | 654 | ------------------------------------------------------------------------------ 655 | -- FIXME: Bundles 656 | 657 | ------------------------------------------------------------------------------ 658 | -- Properties of embedding rationals 659 | 660 | rational-mono : ∀ {q r} → q ℚ.≤ r → rational q ≤ rational r 661 | rational-mono q≤r .*≤* = ℚ.≤-trans q≤r 662 | 663 | rational-cong : ∀ {q r} → q ℚ.≃ r → rational q ≃ rational r 664 | rational-cong q≃r .proj₁ = rational-mono (ℚ.≤-reflexive q≃r) 665 | rational-cong q≃r .proj₂ = rational-mono (ℚ.≤-reflexive (ℚ.≃-sym q≃r)) 666 | 667 | rational+ : ℚ⁺ → ℝᵘ 668 | rational+ q = rational (ℚ⁺.fog q) 669 | 670 | rational-mono-reflect : ∀ {q r} → rational q ≤ rational+ r → q ℚ.≤ ℚ⁺.fog r 671 | rational-mono-reflect {q}{r} q≤r = q≤r .*≤* {r} ℚ.≤-refl 672 | 673 | rational-0 : rational 0ℚ ≃ 0ℝ 674 | rational-0 .proj₁ .*≤* {q} tt = ℚ⁺.fog-nonneg q 675 | rational-0 .proj₂ = 0-least _ 676 | 677 | rational-+ : ∀ q r → 0ℚ ℚ.≤ q → 0ℚ ℚ.≤ r → (rational q + rational r) ≃ rational (q ℚ.+ r) 678 | rational-+ q r 0≤q 0≤r .proj₁ .*≤* {ε} q+r≤ε s = 679 | ℚ⁺.nn+pos q (s /2) 0≤q , 680 | ℚ⁺.nn+pos r (s /2) 0≤r , 681 | ℚ⁺.r≤r (begin 682 | q ℚ.+ ℚ⁺.fog (s /2) ℚ.+ (r ℚ.+ ℚ⁺.fog (s /2)) 683 | ≈⟨ ℚ-interchange q (ℚ⁺.fog (s /2)) r (ℚ⁺.fog (s /2)) ⟩ 684 | (q ℚ.+ r) ℚ.+ (ℚ⁺.fog (s /2) ℚ.+ ℚ⁺.fog (s /2)) 685 | ≈⟨ ℚ.+-congʳ (q ℚ.+ r) (ℚ.≃-sym (ℚ⁺.+-fog (s /2) (s /2))) ⟩ 686 | (q ℚ.+ r) ℚ.+ ℚ⁺.fog (s /2 ℚ⁺.+ s /2) 687 | ≤⟨ ℚ.+-mono-≤ q+r≤ε (ℚ⁺.fog-mono (ℚ⁺.≤-reflexive (ℚ⁺.half+half {s}))) ⟩ 688 | ℚ⁺.fog ε ℚ.+ ℚ⁺.fog s 689 | ≈⟨ ℚ⁺.+-fog ε s ⟩ 690 | ℚ⁺.fog (ε ℚ⁺.+ s) 691 | ∎) , 692 | ℚ⁺.q≤nn+pos q (s /2) , 693 | ℚ⁺.q≤nn+pos r (s /2) 694 | where open ℚ.≤-Reasoning 695 | rational-+ q r 0≤q 0≤r .proj₂ .*≤* {ε} q+r∋ε = 696 | rational (q ℚ.+ r) .closed {ε} λ s → 697 | let ε₁ , ε₂ , ε₁+ε₂≤ε+s , q≤ε₁ , r≤ε₂ = q+r∋ε s in 698 | begin 699 | q ℚ.+ r 700 | ≤⟨ ℚ.+-mono-≤ q≤ε₁ r≤ε₂ ⟩ 701 | ℚ⁺.fog ε₁ ℚ.+ ℚ⁺.fog ε₂ 702 | ≈⟨ ℚ.≃-sym (ℚ⁺.+-fog ε₁ ε₂) ⟩ 703 | ℚ⁺.fog (ε₁ ℚ⁺.+ ε₂) 704 | ≤⟨ ℚ⁺.fog-mono ε₁+ε₂≤ε+s ⟩ 705 | ℚ⁺.fog (ε ℚ⁺.+ s) 706 | ∎ 707 | where open ℚ.≤-Reasoning 708 | 709 | rational⁺-* : ∀ q r → (rational+ q * rational+ r) ≃ rational+ (q ℚ⁺.* r) 710 | rational⁺-* q r .proj₁ .*≤* {ε} qr≤ε s = 711 | q , 712 | r , 713 | (begin 714 | q ℚ⁺.* r 715 | ≤⟨ ℚ⁺.r≤r qr≤ε ⟩ 716 | ε 717 | ≤⟨ ℚ⁺.+-increasing ⟩ 718 | ε ℚ⁺.+ s 719 | ∎) , 720 | ℚ.≤-refl , 721 | ℚ.≤-refl 722 | where open ℚ⁺.≤-Reasoning 723 | rational⁺-* q r .proj₂ .*≤* {ε} qr∋ε = 724 | rational+ (q ℚ⁺.* r) .closed {ε} λ s → 725 | let ε₁ , ε₂ , ε₁ε₂≤ε+s , q≤ε₁ , r≤ε₂ = qr∋ε s in 726 | begin 727 | ℚ⁺.fog q ℚ.* ℚ⁺.fog r 728 | ≤⟨ ℚ.*-monoʳ-≤-pos {ℚ⁺.fog q} (ℚ.positive (ℚ⁺.fog-positive q)) r≤ε₂ ⟩ 729 | ℚ⁺.fog q ℚ.* ℚ⁺.fog ε₂ 730 | ≤⟨ ℚ.*-monoˡ-≤-pos (ℚ.positive (ℚ⁺.fog-positive ε₂)) q≤ε₁ ⟩ 731 | ℚ⁺.fog ε₁ ℚ.* ℚ⁺.fog ε₂ 732 | ≈⟨ ℚ.≃-sym (ℚ⁺.*-fog ε₁ ε₂) ⟩ 733 | ℚ⁺.fog (ε₁ ℚ⁺.* ε₂) 734 | ≤⟨ ℚ⁺.fog-mono ε₁ε₂≤ε+s ⟩ 735 | ℚ⁺.fog (ε ℚ⁺.+ s) 736 | ∎ 737 | where open ℚ.≤-Reasoning 738 | 739 | rational⁺-+ : ∀ q r → (rational+ q + rational+ r) ≃ rational+ (q ℚ⁺.+ r) 740 | rational⁺-+ q r = 741 | begin-equality 742 | rational+ q + rational+ r 743 | ≡⟨⟩ 744 | rational (ℚ⁺.fog q) + rational (ℚ⁺.fog r) 745 | ≈⟨ rational-+ (ℚ⁺.fog q) (ℚ⁺.fog r) (ℚ⁺.fog-nonneg q) (ℚ⁺.fog-nonneg r) ⟩ 746 | rational (ℚ⁺.fog q ℚ.+ ℚ⁺.fog r) 747 | ≈⟨ rational-cong (ℚ⁺.+-fog q r) ⟩ 748 | rational+ (q ℚ⁺.+ r) 749 | ∎ 750 | where open ≤-Reasoning 751 | 752 | rational+<∞ : ∀ q → rational+ q < ∞ 753 | rational+<∞ q = q , ℚ.≤-refl , (λ x → x) 754 | 755 | module _ where 756 | open import Data.Integer 757 | open import Data.Nat 758 | 759 | 0≤1 : 0ℚ ℚ.≤ 1ℚ 760 | 0≤1 = ℚ.*≤* (+≤+ z≤n) 761 | 762 | rational<∞ : ∀ q → rational q < ∞ 763 | rational<∞ q with ℚ.<-cmp 0ℚ q 764 | ... | Relation.Binary.tri< a ¬b ¬c = rational+<∞ ℚ⁺.⟨ q , (ℚ.positive a) ⟩ 765 | ... | Relation.Binary.tri≈ ¬a b ¬c = ℚ⁺.1ℚ⁺ , ℚ.≤-trans (ℚ.≤-reflexive (ℚ.≃-sym b)) 0≤1 , λ x → x 766 | ... | Relation.Binary.tri> ¬a ¬b c = ℚ⁺.1ℚ⁺ , ℚ.<⇒≤ (ℚ.<-≤-trans c 0≤1) , λ x → x 767 | 768 | rational-* : ∀ q r → 0ℚ ℚ.≤ q → 0ℚ ℚ.≤ r → (rational q * rational r) ≃ rational (q ℚ.* r) 769 | rational-* q r 0≤q 0≤r with ℚ.<-cmp 0ℚ q 770 | ... | Relation.Binary.tri≈ _ q≃0 _ = 771 | begin-equality 772 | rational q * rational r 773 | ≈⟨ *-cong (rational-cong (ℚ.≃-sym q≃0)) ≃-refl ⟩ 774 | rational 0ℚ * rational r 775 | ≈⟨ *-cong rational-0 ≃-refl ⟩ 776 | 0ℝ * rational r 777 | ≈⟨ *-zeroˡ (rational r) (rational<∞ r) ⟩ 778 | 0ℝ 779 | ≈⟨ ≃-sym rational-0 ⟩ 780 | rational 0ℚ 781 | ≈⟨ rational-cong (ℚ.≃-sym (ℚ.*-zeroˡ r)) ⟩ 782 | rational (0ℚ ℚ.* r) 783 | ≈⟨ rational-cong (ℚ.*-cong q≃0 (ℚ.≃-refl {r})) ⟩ 784 | rational (q ℚ.* r) 785 | ∎ 786 | where open ≤-Reasoning 787 | ... | Relation.Binary.tri> _ _ q<0 = Data.Empty.⊥-elim (ℚ.<⇒≱ q<0 0≤q) 788 | ... | Relation.Binary.tri< 0 _ _ r<0 = Data.Empty.⊥-elim (ℚ.<⇒≱ r<0 0≤r) 807 | ... | Relation.Binary.tri< 0) 24 | 25 | 26 | ------------------------------------------------------------------------------ 27 | -- the real numbers as the metric completion of the rationals 28 | 29 | ℝspc : MSpc 30 | ℝspc = 𝒞 ℚspc 31 | 32 | ℝspc[_] : ℚ⁺ → MSpc 33 | ℝspc[ b ] = 𝒞 ℚspc[ b ] 34 | 35 | ℝspc[_⟩ : ℚ⁺ → MSpc 36 | ℝspc[ a ⟩ = 𝒞 ℚspc[ a ⟩ 37 | 38 | ℝ : Set 39 | ℝ = ℝspc .MSpc.Carrier 40 | 41 | ℝ[_] : ℚ⁺ → Set 42 | ℝ[ q ] = ℝspc[ q ] .MSpc.Carrier 43 | 44 | ------------------------------------------------------------------------------ 45 | 46 | _≃_ : ℝ → ℝ → Set 47 | x ≃ y = MSpc._≈_ ℝspc x y 48 | 49 | ≃-refl : ∀ {x} → x ≃ x 50 | ≃-refl {x} = MSpc.≈-refl ℝspc {x} 51 | 52 | ≃-sym : ∀ {x y} → x ≃ y → y ≃ x 53 | ≃-sym {x}{y} = MSpc.≈-sym ℝspc {x} {y} 54 | 55 | ≃-trans : ∀ {x y z} → x ≃ y → y ≃ z → x ≃ z 56 | ≃-trans {x}{y}{z} = MSpc.≈-trans ℝspc {x} {y} {z} 57 | 58 | ≃-isEquivalence : IsEquivalence _≃_ 59 | ≃-isEquivalence .IsEquivalence.refl {x} = ≃-refl {x} 60 | ≃-isEquivalence .IsEquivalence.sym {x} {y} = ≃-sym {x} {y} 61 | ≃-isEquivalence .IsEquivalence.trans {x} {y} {z} = ≃-trans {x} {y} {z} 62 | 63 | ≃-setoid : Setoid 0ℓ 0ℓ 64 | ≃-setoid .Setoid.Carrier = ℝ 65 | ≃-setoid .Setoid._≈_ = _≃_ 66 | ≃-setoid .Setoid.isEquivalence = ≃-isEquivalence 67 | 68 | ------------------------------------------------------------------------------ 69 | -- Arithmetic as morphisms in Metric Spaces 70 | 71 | -- FIXME: this seems to type check really slowly without no-eta-equality on MSpc and _⇒_ 72 | addℝ : (ℝspc ⊗ ℝspc) ⇒ ℝspc 73 | addℝ = 𝒞-map add ∘ monoidal-⊗ 74 | 75 | negateℝ : ℝspc ⇒ ℝspc 76 | negateℝ = 𝒞-map negate 77 | 78 | zeroℝ : ⊤ₘ ⇒ ℝspc 79 | zeroℝ = 𝒞-unit ∘ zero 80 | 81 | -- FIXME: now to prove that this gives an abelian group. This is true 82 | -- for any monoid + monoidal functor, and should probably rely on 83 | -- results from the agda-categories library. AFAICT (2021-12-01) I 84 | -- don't think this result is in the agda-categories library. 85 | addℝ-comm : (addℝ ∘ ⊗-symmetry) ≈f addℝ 86 | addℝ-comm = 87 | begin 88 | addℝ ∘ ⊗-symmetry 89 | ≡⟨⟩ 90 | (𝒞-map add ∘ monoidal-⊗) ∘ ⊗-symmetry 91 | ≈˘⟨ assoc (𝒞-map add) monoidal-⊗ ⊗-symmetry ⟩ 92 | 𝒞-map add ∘ (monoidal-⊗ ∘ ⊗-symmetry) 93 | ≈⟨ ∘-cong (≈f-isEquivalence .IsEquivalence.refl) monoidal-symmetry ⟩ 94 | 𝒞-map add ∘ (𝒞-map ⊗-symmetry ∘ monoidal-⊗) 95 | ≈⟨ assoc (𝒞-map add) (𝒞-map ⊗-symmetry) monoidal-⊗ ⟩ 96 | (𝒞-map add ∘ 𝒞-map ⊗-symmetry) ∘ monoidal-⊗ 97 | ≈˘⟨ ∘-cong 𝒞-map-∘ (≈f-isEquivalence .IsEquivalence.refl) ⟩ 98 | 𝒞-map (add ∘ ⊗-symmetry) ∘ monoidal-⊗ 99 | ≈⟨ ∘-cong (𝒞-map-cong add-comm) (≈f-isEquivalence .IsEquivalence.refl) ⟩ 100 | 𝒞-map add ∘ monoidal-⊗ 101 | ≡⟨⟩ 102 | addℝ 103 | ∎ 104 | where open import Relation.Binary.Reasoning.Setoid ≈f-setoid 105 | open import Relation.Binary using (IsEquivalence; Setoid) 106 | 107 | addℝ-assoc : (addℝ ∘ (addℝ ⊗f id) ∘ ⊗-assoc) ≈f (addℝ ∘ (id ⊗f addℝ)) 108 | addℝ-assoc = 109 | begin 110 | addℝ ∘ (addℝ ⊗f id) ∘ ⊗-assoc 111 | ≡⟨⟩ 112 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map add ∘ monoidal-⊗) ⊗f id) ∘ ⊗-assoc 113 | ≈˘⟨ ∘-cong refl (∘-cong (⊗f-cong refl (identityˡ id)) refl) ⟩ 114 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map add ∘ monoidal-⊗) ⊗f (id ∘ id)) ∘ ⊗-assoc 115 | ≈˘⟨ ∘-cong refl (∘-cong (⊗f-cong refl (∘-cong 𝒞-map-id refl)) refl) ⟩ 116 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map add ∘ monoidal-⊗) ⊗f (𝒞-map id ∘ id)) ∘ ⊗-assoc 117 | ≈⟨ ∘-cong refl (∘-cong (⊗f-∘ (𝒞-map add) monoidal-⊗ (𝒞-map id) id) refl) ⟩ 118 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map add ⊗f 𝒞-map id) ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 119 | ≈⟨ assoc (𝒞-map add ∘ monoidal-⊗) ((𝒞-map add ⊗f 𝒞-map id) ∘ (monoidal-⊗ ⊗f id)) ⊗-assoc ⟩ 120 | ((𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map add ⊗f 𝒞-map id) ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 121 | ≈⟨ ∘-cong (assoc _ _ _) refl ⟩ 122 | (((𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map add ⊗f 𝒞-map id)) ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 123 | ≈˘⟨ ∘-cong (∘-cong (assoc _ _ _) refl) refl ⟩ 124 | ((𝒞-map add ∘ monoidal-⊗ ∘ (𝒞-map add ⊗f 𝒞-map id)) ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 125 | ≈⟨ ∘-cong (∘-cong (∘-cong refl (monoidal-natural add id)) refl) refl ⟩ 126 | ((𝒞-map add ∘ 𝒞-map (add ⊗f id) ∘ monoidal-⊗) ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 127 | ≈⟨ ∘-cong (∘-cong (assoc _ _ _) refl) refl ⟩ 128 | (((𝒞-map add ∘ 𝒞-map (add ⊗f id)) ∘ monoidal-⊗) ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 129 | ≈⟨ ∘-cong (sym (assoc _ _ _)) refl ⟩ 130 | ((𝒞-map add ∘ 𝒞-map (add ⊗f id)) ∘ monoidal-⊗ ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 131 | ≈⟨ sym (assoc _ _ _) ⟩ 132 | (𝒞-map add ∘ 𝒞-map (add ⊗f id)) ∘ (monoidal-⊗ ∘ (monoidal-⊗ ⊗f id)) ∘ ⊗-assoc 133 | ≈⟨ ∘-cong (sym 𝒞-map-∘) (sym (assoc _ _ _)) ⟩ 134 | 𝒞-map (add ∘ (add ⊗f id)) ∘ (monoidal-⊗ ∘ (monoidal-⊗ ⊗f id) ∘ ⊗-assoc) 135 | ≈⟨ ∘-cong refl monoidal-assoc ⟩ 136 | 𝒞-map (add ∘ (add ⊗f id)) ∘ 𝒞-map ⊗-assoc ∘ monoidal-⊗ ∘ (id ⊗f monoidal-⊗) 137 | ≈⟨ assoc _ _ _ ⟩ 138 | (𝒞-map (add ∘ (add ⊗f id)) ∘ 𝒞-map ⊗-assoc) ∘ monoidal-⊗ ∘ (id ⊗f monoidal-⊗) 139 | ≈⟨ ∘-cong (sym 𝒞-map-∘) refl ⟩ 140 | 𝒞-map ((add ∘ (add ⊗f id)) ∘ ⊗-assoc) ∘ monoidal-⊗ ∘ (id ⊗f monoidal-⊗) 141 | ≈⟨ ∘-cong (𝒞-map-cong (sym (assoc _ _ _))) refl ⟩ 142 | 𝒞-map (add ∘ (add ⊗f id) ∘ ⊗-assoc) ∘ monoidal-⊗ ∘ (id ⊗f monoidal-⊗) 143 | ≈⟨ ∘-cong (𝒞-map-cong add-assoc) refl ⟩ 144 | 𝒞-map (add ∘ (id ⊗f add)) ∘ monoidal-⊗ ∘ (id ⊗f monoidal-⊗) 145 | ≈⟨ ∘-cong 𝒞-map-∘ refl ⟩ 146 | (𝒞-map add ∘ 𝒞-map (id ⊗f add)) ∘ monoidal-⊗ ∘ (id ⊗f monoidal-⊗) 147 | ≈⟨ assoc _ _ _ ⟩ 148 | ((𝒞-map add ∘ 𝒞-map (id ⊗f add)) ∘ monoidal-⊗) ∘ (id ⊗f monoidal-⊗) 149 | ≈⟨ ∘-cong (sym (assoc _ _ _)) refl ⟩ 150 | (𝒞-map add ∘ 𝒞-map (id ⊗f add) ∘ monoidal-⊗) ∘ (id ⊗f monoidal-⊗) 151 | ≈⟨ ∘-cong (∘-cong refl (sym (monoidal-natural id add))) refl ⟩ 152 | (𝒞-map add ∘ monoidal-⊗ ∘ (𝒞-map id ⊗f 𝒞-map add)) ∘ (id ⊗f monoidal-⊗) 153 | ≈⟨ ∘-cong (assoc _ _ _) refl ⟩ 154 | ((𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map id ⊗f 𝒞-map add)) ∘ (id ⊗f monoidal-⊗) 155 | ≈⟨ sym (assoc _ _ _) ⟩ 156 | (𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map id ⊗f 𝒞-map add) ∘ (id ⊗f monoidal-⊗) 157 | ≈⟨ ∘-cong refl (sym (⊗f-∘ (𝒞-map id) id (𝒞-map add) monoidal-⊗)) ⟩ 158 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map id ∘ id) ⊗f (𝒞-map add ∘ monoidal-⊗)) 159 | ≈⟨ ∘-cong refl (⊗f-cong (∘-cong 𝒞-map-id refl) refl) ⟩ 160 | (𝒞-map add ∘ monoidal-⊗) ∘ ((id ∘ id) ⊗f (𝒞-map add ∘ monoidal-⊗)) 161 | ≈⟨ ∘-cong refl (⊗f-cong (identityˡ id) refl) ⟩ 162 | (𝒞-map add ∘ monoidal-⊗) ∘ (id ⊗f (𝒞-map add ∘ monoidal-⊗)) 163 | ≡⟨⟩ 164 | addℝ ∘ (id ⊗f addℝ) 165 | ∎ 166 | where open import Relation.Binary.Reasoning.Setoid ≈f-setoid 167 | open import Relation.Binary using (IsEquivalence; Setoid) 168 | refl : ∀ {X Y} {f : X ⇒ Y} → f ≈f f 169 | refl = ≈f-isEquivalence .IsEquivalence.refl 170 | sym : ∀ {X Y} {f g : X ⇒ Y} → f ≈f g → g ≈f f 171 | sym = ≈f-isEquivalence .IsEquivalence.sym 172 | 173 | -- This proof is very simple, but very long! 174 | addℝ-identityˡ : (addℝ ∘ (zeroℝ ⊗f id) ∘ left-unit⁻¹) ≈f id 175 | addℝ-identityˡ = 176 | begin 177 | addℝ ∘ (zeroℝ ⊗f id) ∘ left-unit⁻¹ 178 | ≡⟨⟩ 179 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-unit ∘ zero) ⊗f id) ∘ left-unit⁻¹ 180 | ≈⟨ ∘-cong refl (∘-cong (⊗f-cong unit-natural refl) refl) ⟩ 181 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map zero ∘ 𝒞-unit) ⊗f id) ∘ left-unit⁻¹ 182 | ≈˘⟨ ∘-cong refl (∘-cong (⊗f-cong refl (identityˡ id)) refl) ⟩ 183 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map zero ∘ 𝒞-unit) ⊗f (id ∘ id)) ∘ left-unit⁻¹ 184 | ≈˘⟨ ∘-cong refl (∘-cong (⊗f-cong refl (∘-cong 𝒞-map-id refl)) refl) ⟩ 185 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map zero ∘ 𝒞-unit) ⊗f (𝒞-map id ∘ id)) ∘ left-unit⁻¹ 186 | ≈⟨ ∘-cong refl (∘-cong (⊗f-∘ (𝒞-map zero) 𝒞-unit (𝒞-map id) id) refl) ⟩ 187 | (𝒞-map add ∘ monoidal-⊗) ∘ ((𝒞-map zero ⊗f 𝒞-map id) ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 188 | ≈⟨ assoc (𝒞-map add ∘ monoidal-⊗) ((𝒞-map zero ⊗f 𝒞-map id) ∘ (𝒞-unit ⊗f id)) left-unit⁻¹ ⟩ 189 | ((𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map zero ⊗f 𝒞-map id) ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 190 | ≈⟨ ∘-cong (assoc (𝒞-map add ∘ monoidal-⊗) (𝒞-map zero ⊗f 𝒞-map id) (𝒞-unit ⊗f id)) refl ⟩ 191 | (((𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map zero ⊗f 𝒞-map id)) ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 192 | ≈˘⟨ ∘-cong (∘-cong (assoc (𝒞-map add) monoidal-⊗ (𝒞-map zero ⊗f 𝒞-map id)) refl) refl ⟩ 193 | ((𝒞-map add ∘ monoidal-⊗ ∘ (𝒞-map zero ⊗f 𝒞-map id)) ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 194 | ≈⟨ ∘-cong (∘-cong (∘-cong refl (monoidal-natural zero id)) refl) refl ⟩ 195 | ((𝒞-map add ∘ 𝒞-map (zero ⊗f id) ∘ monoidal-⊗) ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 196 | ≈⟨ ∘-cong (∘-cong (assoc (𝒞-map add) (𝒞-map (zero ⊗f id)) monoidal-⊗) refl) refl ⟩ 197 | (((𝒞-map add ∘ 𝒞-map (zero ⊗f id)) ∘ monoidal-⊗) ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 198 | ≈˘⟨ ∘-cong (∘-cong (∘-cong 𝒞-map-∘ refl) refl) refl ⟩ 199 | ((𝒞-map (add ∘ (zero ⊗f id)) ∘ monoidal-⊗) ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 200 | ≈˘⟨ ∘-cong (assoc (𝒞-map (add ∘ (zero ⊗f id))) monoidal-⊗ (𝒞-unit ⊗f id)) refl ⟩ 201 | (𝒞-map (add ∘ (zero ⊗f id)) ∘ monoidal-⊗ ∘ (𝒞-unit ⊗f id)) ∘ left-unit⁻¹ 202 | ≈⟨ ∘-cong (∘-cong refl monoidal-left-unit) refl ⟩ 203 | (𝒞-map (add ∘ (zero ⊗f id)) ∘ 𝒞-map left-unit⁻¹ ∘ left-unit) ∘ left-unit⁻¹ 204 | ≈⟨ ∘-cong (assoc (𝒞-map (add ∘ (zero ⊗f id))) (𝒞-map left-unit⁻¹) left-unit) refl ⟩ 205 | ((𝒞-map (add ∘ (zero ⊗f id)) ∘ 𝒞-map left-unit⁻¹) ∘ left-unit) ∘ left-unit⁻¹ 206 | ≈˘⟨ ∘-cong (∘-cong 𝒞-map-∘ refl) refl ⟩ 207 | (𝒞-map ((add ∘ (zero ⊗f id)) ∘ left-unit⁻¹) ∘ left-unit) ∘ left-unit⁻¹ 208 | ≈˘⟨ assoc (𝒞-map ((add ∘ (zero ⊗f id)) ∘ left-unit⁻¹)) left-unit left-unit⁻¹ ⟩ 209 | 𝒞-map ((add ∘ (zero ⊗f id)) ∘ left-unit⁻¹) ∘ left-unit ∘ left-unit⁻¹ 210 | ≈⟨ ∘-cong (𝒞-map-cong (sym (assoc add (zero ⊗f id) left-unit⁻¹))) left-unit-iso₁ ⟩ 211 | 𝒞-map (add ∘ (zero ⊗f id) ∘ left-unit⁻¹) ∘ id 212 | ≈⟨ ∘-cong (𝒞-map-cong add-identityˡ) refl ⟩ 213 | 𝒞-map id ∘ id 214 | ≈⟨ ∘-cong 𝒞-map-id refl ⟩ 215 | id ∘ id 216 | ≈⟨ identityˡ id ⟩ 217 | id 218 | ∎ 219 | where open import Relation.Binary.Reasoning.Setoid ≈f-setoid 220 | open import Relation.Binary using (IsEquivalence; Setoid) 221 | refl : ∀ {X Y} {f : X ⇒ Y} → f ≈f f 222 | refl = ≈f-isEquivalence .IsEquivalence.refl 223 | sym : ∀ {X Y} {f g : X ⇒ Y} → f ≈f g → g ≈f f 224 | sym = ≈f-isEquivalence .IsEquivalence.sym 225 | 226 | addℝ-inverse : (addℝ ∘ (id ⊗f negateℝ) ∘ (derelict ⊗f derelict) ∘ duplicate) ≈f (zeroℝ ∘ discard) 227 | addℝ-inverse = 228 | begin 229 | addℝ ∘ (id ⊗f negateℝ) ∘ (derelict ⊗f derelict) ∘ duplicate 230 | ≡⟨⟩ 231 | (𝒞-map add ∘ monoidal-⊗) ∘ (id ⊗f 𝒞-map negate) ∘ (derelict ⊗f derelict) ∘ duplicate 232 | ≈⟨ ∘-cong refl (∘-cong (⊗f-cong (sym 𝒞-map-id) refl) refl) ⟩ 233 | (𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map id ⊗f 𝒞-map negate) ∘ (derelict ⊗f derelict) ∘ duplicate 234 | ≈⟨ assoc _ _ _ ⟩ 235 | ((𝒞-map add ∘ monoidal-⊗) ∘ (𝒞-map id ⊗f 𝒞-map negate)) ∘ (derelict ⊗f derelict) ∘ duplicate 236 | ≈⟨ ∘-cong (sym (assoc _ _ _)) refl ⟩ 237 | (𝒞-map add ∘ monoidal-⊗ ∘ (𝒞-map id ⊗f 𝒞-map negate)) ∘ (derelict ⊗f derelict) ∘ duplicate 238 | ≈⟨ ∘-cong (∘-cong refl (monoidal-natural id negate)) refl ⟩ 239 | (𝒞-map add ∘ 𝒞-map (id ⊗f negate) ∘ monoidal-⊗) ∘ (derelict ⊗f derelict) ∘ duplicate 240 | ≈⟨ {!!} ⟩ -- FIXME: need to prove more properties of ![_] and its interaction with 𝒞 241 | zeroℝ ∘ discard 242 | ∎ 243 | where open import Relation.Binary.Reasoning.Setoid ≈f-setoid 244 | open import Relation.Binary using (IsEquivalence; Setoid) 245 | refl : ∀ {X Y} {f : X ⇒ Y} → f ≈f f 246 | refl = ≈f-isEquivalence .IsEquivalence.refl 247 | sym : ∀ {X Y} {f g : X ⇒ Y} → f ≈f g → g ≈f f 248 | sym = ≈f-isEquivalence .IsEquivalence.sym 249 | 250 | 251 | ------------------------------------------------------------------------------ 252 | -- Arithmetic operations as plain functions 253 | 254 | -- FIXME: rename these to remove the ℝ suffixes 255 | 256 | ℚ→ℝ : ℚ → ℝ 257 | ℚ→ℝ q = 𝒞-unit ._⇒_.fun q 258 | 259 | _+ℝ_ : ℝ → ℝ → ℝ 260 | x +ℝ y = addℝ ._⇒_.fun (x , y) 261 | 262 | -ℝ_ : ℝ → ℝ 263 | -ℝ_ x = negateℝ ._⇒_.fun x 264 | 265 | 0ℝ : ℝ 266 | 0ℝ = zeroℝ ._⇒_.fun tt 267 | 268 | _-ℝ_ : ℝ → ℝ → ℝ 269 | x -ℝ y = x +ℝ (-ℝ y) 270 | 271 | ------------------------------------------------------------------------------ 272 | -- Properties of _+_ and -_ 273 | 274 | +ℝ-cong : Congruent₂ _≃_ _+ℝ_ -- ∀ {x₁ x₂ y₁ y₂} → x₁ ≃ x₂ → y₁ ≃ y₂ → (x₁ +ℝ y₁) ≃ (x₂ +ℝ y₂) 275 | +ℝ-cong {x₁}{x₂}{y₁}{y₂} x₁≃x₂ y₁≃y₂ = 276 | _⇒_.cong addℝ {x₁ , y₁} {x₂ , y₂} 277 | (≈-⊗ {ℝspc}{ℝspc} {x₁}{x₂}{y₁}{y₂} x₁≃x₂ y₁≃y₂) 278 | 279 | -ℝ-cong : Congruent₁ _≃_ (-ℝ_) 280 | -ℝ-cong {x₁}{x₂} x₁≃x₂ = 281 | _⇒_.cong negateℝ {x₁} {x₂} x₁≃x₂ 282 | 283 | +ℝ-assoc : ∀ x y z → ((x +ℝ y) +ℝ z) ≃ (x +ℝ (y +ℝ z)) 284 | +ℝ-assoc x y z = addℝ-assoc ._≈f_.f≈f (x , (y , z)) 285 | 286 | +ℝ-comm : ∀ x y → (x +ℝ y) ≃ (y +ℝ x) 287 | +ℝ-comm x y = addℝ-comm ._≈f_.f≈f (y , x) 288 | 289 | +ℝ-identityˡ : ∀ x → (0ℝ +ℝ x) ≃ x 290 | +ℝ-identityˡ x = addℝ-identityˡ ._≈f_.f≈f x 291 | 292 | +ℝ-identityʳ : ∀ x → (x +ℝ 0ℝ) ≃ x 293 | +ℝ-identityʳ x = 294 | begin 295 | x +ℝ 0ℝ 296 | ≈⟨ +ℝ-comm x 0ℝ ⟩ 297 | 0ℝ +ℝ x 298 | ≈⟨ +ℝ-identityˡ x ⟩ 299 | x 300 | ∎ 301 | where open import Relation.Binary.Reasoning.Setoid ≃-setoid 302 | 303 | +-identity : Identity _≃_ 0ℝ _+ℝ_ 304 | +-identity = +ℝ-identityˡ , +ℝ-identityʳ 305 | 306 | +-inverseʳ : RightInverse _≃_ 0ℝ -ℝ_ _+ℝ_ 307 | +-inverseʳ x = addℝ-inverse ._≈f_.f≈f x 308 | 309 | +-inverseˡ : LeftInverse _≃_ 0ℝ -ℝ_ _+ℝ_ 310 | +-inverseˡ x = 311 | begin 312 | (-ℝ x) +ℝ x ≈⟨ +ℝ-comm (-ℝ x) x ⟩ 313 | x +ℝ (-ℝ x) ≈⟨ +-inverseʳ x ⟩ 314 | 0ℝ ∎ 315 | where open import Relation.Binary.Reasoning.Setoid ≃-setoid 316 | 317 | +-inverse : Inverse _≃_ 0ℝ -ℝ_ _+ℝ_ 318 | +-inverse = +-inverseˡ , +-inverseʳ 319 | 320 | ------------------------------------------------------------------------------ 321 | -- Packaging up the proofs of algeraic properties 322 | 323 | +-isMagma : IsMagma _≃_ _+ℝ_ 324 | +-isMagma = record { isEquivalence = ≃-isEquivalence ; ∙-cong = λ {x₁}{x₂}{y₁}{y₂} → +ℝ-cong {x₁}{x₂}{y₁}{y₂} } 325 | 326 | +-isSemigroup : IsSemigroup _≃_ _+ℝ_ 327 | +-isSemigroup = record { isMagma = +-isMagma ; assoc = +ℝ-assoc } 328 | 329 | +-isCommutativeSemigroup : IsCommutativeSemigroup _≃_ _+ℝ_ 330 | +-isCommutativeSemigroup = record { isSemigroup = +-isSemigroup ; comm = +ℝ-comm } 331 | 332 | +-0-isMonoid : IsMonoid _≃_ _+ℝ_ 0ℝ 333 | +-0-isMonoid = record { isSemigroup = +-isSemigroup ; identity = +-identity } 334 | 335 | +-0-isCommutativeMonoid : IsCommutativeMonoid _≃_ _+ℝ_ 0ℝ 336 | +-0-isCommutativeMonoid = record { isMonoid = +-0-isMonoid ; comm = +ℝ-comm } 337 | 338 | +-0-isGroup : IsGroup _≃_ _+ℝ_ 0ℝ (-ℝ_) 339 | +-0-isGroup = record { isMonoid = +-0-isMonoid ; inverse = +-inverse ; ⁻¹-cong = λ {x₁}{x₂} → -ℝ-cong {x₁}{x₂} } 340 | 341 | +-0-isAbelianGroup : IsAbelianGroup _≃_ _+ℝ_ 0ℝ (-ℝ_) 342 | +-0-isAbelianGroup = record { isGroup = +-0-isGroup ; comm = +ℝ-comm } 343 | 344 | +-0-AbelianGroup : AbelianGroup 0ℓ 0ℓ 345 | +-0-AbelianGroup = record { isAbelianGroup = +-0-isAbelianGroup } 346 | 347 | -- FIXME: bundles 348 | 349 | 350 | 351 | ------------------------------------------------------------------------------ 352 | -- Order and apartness 353 | 354 | -- FIXME: this is unfinished 355 | 356 | module _ where 357 | 358 | open RegFun 359 | 360 | 0≤ℝ : ℝ → Set 361 | 0≤ℝ x = ∀ ε → ℚ.- ℚ⁺.fog ε ℚ.≤ x .RegFun.rfun ε 362 | 363 | -- https://github.com/coq-community/corn/blob/master/reals/fast/CRGroupOps.v#L177 364 | -- 0≤ℝ-cong : ∀ {x y} → x ≃ y → 0≤ℝ x → 0≤ℝ y 365 | -- 0≤ℝ-cong {x} {y} x≃y 0≤x ε = {!!} 366 | 367 | -- with ℚ.<-cmp (x .rfun ε) (y .rfun ε) 368 | -- ... | tri< x⟨ε⟩ _ _ y⟨ε⟩ xε - yε ≤ 2ε 398 | -- ==> xε - 2ε ≤ yε 399 | -- 400 | 401 | -- ≈-𝒞 {ℚspc} {x}{y} x≃y ε ε : difference between x(ε) and y(ε) is less than (ε + ε) 402 | -- otherwise, x .rfun ε + (ε + ε) ≤ y .rfun ε 403 | -- so ε ≤ y .rfun ε and so - ε ≤ y .rfun ε 404 | 405 | -- “Not greater than” 406 | _≤ℝ_ : ℝ → ℝ → Set 407 | x ≤ℝ y = 0≤ℝ (y -ℝ x) 408 | 409 | ≤ℝ-refl : ∀ x → x ≤ℝ x 410 | ≤ℝ-refl x ε = {!!} 411 | 412 | 0<ℝ : ℝ → Set 413 | 0<ℝ x = Σ[ ε ∈ ℚ⁺ ] (𝒞-unit ._⇒_.fun (ℚ⁺.fog ε) ≤ℝ x ) 414 | 415 | -- “Greater than” 416 | _<ℝ_ : ℝ → ℝ → Set 417 | x <ℝ y = 0<ℝ (y -ℝ x) 418 | 419 | -- Apartness 420 | _#ℝ_ : ℝ → ℝ → Set 421 | x #ℝ y = x <ℝ y ⊎ y <ℝ x 422 | 423 | -- FIXME: axioms for apartness 424 | 425 | -- FIXME: prove that the distance function on the reals is really the 426 | -- absolute value of the difference. This would require a mapping of 427 | -- the metric space reals into the upper reals. 428 | 429 | {- 430 | module _ where 431 | 432 | open import upper-reals 433 | open upper-reals.ℝᵘ 434 | 435 | to-upper-real : ℝ → ℝᵘ 436 | to-upper-real x .contains q = {!x ≤ℝ (ℝ→ℚ q)!} 437 | to-upper-real x .upper = {!!} 438 | to-upper-real x .closed = {!!} 439 | -} 440 | 441 | ------------------------------------------------------------------------------ 442 | -- Multiplication and reciprocal 443 | 444 | -- scaling a real by a positive rational 445 | scale : ∀ q → ![ q ] ℝspc ⇒ ℝspc 446 | scale q = 𝒞-map (ℚ-scale q) ∘ distr q 447 | 448 | -- Fully "metrised" versions of multiplication and reciprocal 449 | 450 | mulℝ : ∀ a b → (![ b ] ℝspc[ a ] ⊗ ![ a ] ℝspc[ b ]) ⇒ ℝspc[ a ℚ⁺.* b ] 451 | mulℝ a b = 𝒞-map (mul a b) ∘ monoidal-⊗ ∘ (distr b ⊗f distr a) 452 | 453 | reciprocalℝ : ∀ a → ![ ℚ⁺.1/ (a ℚ⁺.* a) ] ℝspc[ a ⟩ ⇒ ℝspc 454 | reciprocalℝ a = 𝒞-map (recip a) ∘ distr (ℚ⁺.1/ (a ℚ⁺.* a)) 455 | 456 | ------------------------------------------------------------------------------ 457 | -- Bounding regular functions 458 | 459 | open import Data.Product 460 | import Data.Rational.Unnormalised as ℚ 461 | import Data.Rational.Unnormalised.Properties as ℚ 462 | 463 | module _ where 464 | import Data.Integer as ℤ 465 | import Data.Nat as ℕ 466 | 467 | 0<½ : ℚ.0ℚᵘ ℚ.< ℚ.½ 468 | 0<½ = ℚ.*<* (ℤ.+<+ (ℕ.s≤s ℕ.z≤n)) 469 | 470 | get-bound : ℝ → ℚ⁺ 471 | get-bound r = 472 | ℚ⁺.⟨ (ℚ.∣ r .RegFun.rfun ℚ⁺.½ ∣ ℚ.+ ℚ.½) 473 | , ℚ.positive 474 | (begin-strict 475 | ℚ.0ℚᵘ 476 | ≤⟨ ℚ.0≤∣p∣ (r .RegFun.rfun ℚ⁺.½) ⟩ 477 | ℚ.∣ r .RegFun.rfun ℚ⁺.½ ∣ 478 | ≈⟨ ℚ.≃-sym (ℚ.+-identityʳ ℚ.∣ r .RegFun.rfun ℚ⁺.½ ∣) ⟩ 479 | ℚ.∣ r .RegFun.rfun ℚ⁺.½ ∣ ℚ.+ ℚ.0ℚᵘ 480 | <⟨ ℚ.+-mono-≤-< (ℚ.≤-refl {ℚ.∣ r .RegFun.rfun ℚ⁺.½ ∣}) 0<½ ⟩ 481 | ℚ.∣ r .RegFun.rfun ℚ⁺.½ ∣ ℚ.+ ℚ.½ 482 | ∎) 483 | ⟩ 484 | where open ℚ.≤-Reasoning 485 | 486 | bound : ℝ → Σ[ q ∈ ℚ⁺ ] ℝ[ q ] 487 | bound r .proj₁ = get-bound r 488 | bound r .proj₂ = {!!} -- 𝒞-map (clamping.clamp (get-bound r)) ._⇒_.fun r 489 | 490 | open _⇒_ 491 | 492 | ℝ-unbound : ∀ {q} → ℝ[ q ] → ℝ 493 | ℝ-unbound {q} = 𝒞-map (unbound q) .fun 494 | 495 | bound-eq : ∀ x → ℝ-unbound (bound x .proj₂) ≃ x 496 | bound-eq x = 497 | 𝒞-≈ {x = ℝ-unbound (bound x .proj₂)} {y = x} λ ε₁ ε₂ → 498 | begin 499 | ℝᵘ.rational ℚ.∣ ℝ-unbound (bound x .proj₂) .RegFun.rfun ε₁ ℚ.- x .RegFun.rfun ε₂ ∣ 500 | ≡⟨⟩ 501 | ℝᵘ.rational ℚ.∣ clamping.clamp-v (get-bound x) (x .RegFun.rfun ε₁) ℚ.- x .RegFun.rfun ε₂ ∣ 502 | ≤⟨ {!!} ⟩ 503 | ℝᵘ.rational+ (ε₁ ℚ⁺.+ ε₂) 504 | ∎ 505 | where open ℝᵘ.≤-Reasoning 506 | 507 | _*ℝ_ : ℝ → ℝ → ℝ 508 | _*ℝ_ x y = 509 | let a , x' = bound x in 510 | let b , y' = bound y in 511 | ℝ-unbound (mulℝ a b .fun (x' , y')) 512 | 513 | -- TODO: 514 | -- 1. congruence 515 | -- 2. associativity (needs proofs from rationals) 516 | -- 3. distributivity 517 | -- 4. identities 518 | -- 5. zeros 519 | 520 | 2ℝ : ℝ 521 | 2ℝ = ℚ→ℝ (ℚ.1ℚᵘ ℚ.+ ℚ.1ℚᵘ) 522 | 523 | 4ℝ : ℝ 524 | 4ℝ = 2ℝ *ℝ 2ℝ 525 | 526 | 4ℚ : ℚ 527 | 4ℚ = ℚ.1ℚᵘ ℚ.+ ℚ.1ℚᵘ ℚ.+ ℚ.1ℚᵘ ℚ.+ ℚ.1ℚᵘ 528 | 529 | module _ where 530 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 531 | 532 | _ : 4ℝ .RegFun.rfun ℚ⁺.½ ≡ 4ℚ 533 | _ = refl 534 | 535 | ------------------------------------------------------------------------ 536 | -- TODO: reciprocal 537 | 538 | {- If a real is greater than 0 by the above definition, then we can 539 | use the ε from that to bound it from underneath. Then apply the 540 | reciprocal on the rationals. -} 541 | 542 | ------------------------------------------------------------------------ 543 | -- TODO: Alternating decreasing series 544 | 545 | -- If we have: 546 | -- a series a : ℕ → ℚ 547 | -- limit is 0: ∀ ε → Σ[ n ∈ ℕ ] (∣ a n ∣ ≤ ε) 548 | -- alternating: ??? 549 | -- decreasing: ∀ i → ∣ a (suc i) ∣ < ∣ a i ∣ 550 | 551 | -- define x(ε) = sum(modulus(ε), a) 552 | -- then prove that this is a regular function 553 | 554 | -- and then prove that it gives us the infinite sum?? 555 | -------------------------------------------------------------------------------- /src/MetricSpace/Completion.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module MetricSpace.Completion where 4 | 5 | open import Data.Product using (_×_; _,_; proj₁; proj₂; swap) 6 | import Data.Rational.Unnormalised.Positive as ℚ⁺ 7 | open ℚ⁺ using (ℚ⁺; _/2; 1/_; 1ℚ⁺) 8 | open import MetricSpace 9 | open import Data.Real.UpperClosed 10 | 11 | open MSpc 12 | open _⇒_ 13 | open _≈f_ 14 | 15 | record RegFun (X : MSpc) : Set where 16 | field 17 | rfun : ℚ⁺ → X .Carrier 18 | regular : ∀ ε₁ ε₂ → X .dist (rfun ε₁) (rfun ε₂) ≤ rational+ (ε₁ ℚ⁺.+ ε₂) 19 | open RegFun 20 | 21 | reg-dist : ∀ {X} → RegFun X → RegFun X → ℝᵘ 22 | reg-dist {X} x y = 23 | sup (ℚ⁺ × ℚ⁺) (λ { (ε₁ , ε₂) → X .dist (x .rfun ε₁) (y .rfun ε₂) ⊖ (ε₁ ℚ⁺.+ ε₂) } ) 24 | 25 | 26 | 27 | -- FIXME: some lemmas for dealing with reg-dist, to avoid all the 28 | -- dealing with sup-least and ⊖-iso1/2 below 29 | 30 | 31 | 32 | 33 | 𝒞 : MSpc → MSpc 34 | 𝒞 X .Carrier = RegFun X 35 | 𝒞 X .dist = reg-dist 36 | 𝒞 X .refl {x} = sup-least λ { (ε₁ , ε₂) → ≤-trans (⊖-mono (x .regular ε₁ ε₂) ℚ⁺.≤-refl) (⊖-0 (ε₁ ℚ⁺.+ ε₂)) } 37 | 𝒞 X .sym = sup-mono-≤ swap λ { (ε₁ , ε₂) → ⊖-mono (X .sym) (ℚ⁺.≤-reflexive (ℚ⁺.+-comm ε₂ ε₁)) } 38 | 𝒞 X .triangle {x}{y}{z} = 39 | sup-least λ { (ε₁ , ε₂) → closedness λ ε → 40 | begin 41 | (X .dist (x .rfun ε₁) (z .rfun ε₂) ⊖ (ε₁ ℚ⁺.+ ε₂)) ⊖ ε 42 | ≤⟨ ⊖-mono (⊖-mono (X .triangle) ℚ⁺.≤-refl) (ℚ⁺.≤-reflexive ℚ⁺.half+half) ⟩ 43 | ((X .dist (x .rfun ε₁) (y .rfun (ε /2)) + X .dist (y .rfun (ε /2)) (z .rfun ε₂)) ⊖ (ε₁ ℚ⁺.+ ε₂)) ⊖ (ε /2 ℚ⁺.+ ε /2) 44 | ≈⟨ ⊖-+ ⟩ 45 | ((X .dist (x .rfun ε₁) (y .rfun (ε /2)) + X .dist (y .rfun (ε /2)) (z .rfun ε₂)) ⊖ ((ε₁ ℚ⁺.+ ε₂) ℚ⁺.+ (ε /2 ℚ⁺.+ ε /2))) 46 | ≤⟨ ⊖-mono ≤-refl (ℚ⁺.≤-reflexive (ℚ⁺-interchange ε₁ _ _ _)) ⟩ 47 | ((X .dist (x .rfun ε₁) (y .rfun (ε /2)) + X .dist (y .rfun (ε /2)) (z .rfun ε₂)) ⊖ ((ε₁ ℚ⁺.+ ε /2) ℚ⁺.+ (ε₂ ℚ⁺.+ ε /2))) 48 | ≤⟨ ⊖-mono ≤-refl (ℚ⁺.≤-reflexive (ℚ⁺.+-congʳ (ε₁ ℚ⁺.+ ε /2) (ℚ⁺.+-comm _ _))) ⟩ 49 | ((X .dist (x .rfun ε₁) (y .rfun (ε /2)) + X .dist (y .rfun (ε /2)) (z .rfun ε₂)) ⊖ ((ε₁ ℚ⁺.+ ε /2) ℚ⁺.+ (ε /2 ℚ⁺.+ ε₂))) 50 | ≤⟨ ⊖-+-+ ⟩ 51 | (X .dist (x .rfun ε₁) (y .rfun (ε /2)) ⊖ (ε₁ ℚ⁺.+ ε /2)) + (X .dist (y .rfun (ε /2)) (z .rfun ε₂) ⊖ (ε /2 ℚ⁺.+ ε₂)) 52 | ≤⟨ +-mono-≤ (sup-upper (ε₁ , ε /2)) (sup-upper (ε /2 , ε₂)) ⟩ 53 | reg-dist x y + reg-dist y z 54 | ∎ 55 | } 56 | where open ≤-Reasoning 57 | 58 | -- 𝒞 X .dist x y ≤ X .dist (x .rfun ε) (y .rfun ϵ) 59 | 60 | 𝒞-≈ : ∀ {X} {x y : 𝒞 X .Carrier} → 61 | (∀ ε₁ ε₂ → X .dist (x .rfun ε₁) (y .rfun ε₂) ≤ rational+ (ε₁ ℚ⁺.+ ε₂)) → 62 | _≈_ (𝒞 X) x y 63 | 𝒞-≈ {X}{x}{y} h = sup-least λ { (ε₁ , ε₂) → ⊖-iso2 64 | (begin 65 | X .dist (x .rfun ε₁) (y .rfun ε₂) 66 | ≤⟨ h ε₁ ε₂ ⟩ 67 | rational+ (ε₁ ℚ⁺.+ ε₂) 68 | ≈⟨ ≃-sym (+-identityˡ (rational+ (ε₁ ℚ⁺.+ ε₂))) ⟩ 69 | 0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂) 70 | ∎) } 71 | where open ≤-Reasoning 72 | 73 | ------------------------------------------------------------------------------ 74 | ≈-𝒞 : ∀ {X} {x y : 𝒞 X .Carrier} → 75 | _≈_ (𝒞 X) x y → 76 | ∀ ε₁ ε₂ → X .dist (x .rfun ε₁) (y .rfun ε₂) ≤ rational+ (ε₁ ℚ⁺.+ ε₂) 77 | ≈-𝒞 {X}{x}{y} x≃y ε₁ ε₂ = 78 | ⊖-iso1-0 79 | (begin 80 | X .dist (x .rfun ε₁) (y .rfun ε₂) ⊖ (ε₁ ℚ⁺.+ ε₂) 81 | ≤⟨ sup-upper (ε₁ , ε₂) ⟩ 82 | 𝒞 X .dist x y 83 | ≤⟨ x≃y ⟩ 84 | 0ℝ 85 | ∎) 86 | where open ≤-Reasoning 87 | 88 | ------------------------------------------------------------------------------ 89 | -- Functor operation on morphisms 90 | 91 | -- FIXME: does this also work for uniformly continuous functions? 92 | map : ∀ {X Y} → X ⇒ Y → 𝒞 X ⇒ 𝒞 Y 93 | map f .fun x .rfun ε = f .fun (x .rfun ε) 94 | map f .fun x .regular ε₁ ε₂ = ≤-trans (f .non-expansive) (x .regular ε₁ ε₂) 95 | map f .non-expansive = sup-mono-≤ (λ x → x) λ { (ε₁ , ε₂) → ⊖-mono (f .non-expansive) ℚ⁺.≤-refl } 96 | 97 | map-cong : ∀ {X Y}{f₁ f₂ : X ⇒ Y} → f₁ ≈f f₂ → map f₁ ≈f map f₂ 98 | map-cong {X}{Y}{f₁}{f₂} f₁≈f₂ .f≈f x = 99 | 𝒞-≈ {Y} {map f₁ .fun x} {map f₂ .fun x} λ ε₁ ε₂ → 100 | begin 101 | Y .dist (f₁ .fun (x .rfun ε₁)) (f₂ .fun (x .rfun ε₂)) 102 | ≤⟨ Y .triangle ⟩ 103 | Y .dist (f₁ .fun (x .rfun ε₁)) (f₁ .fun (x .rfun ε₂)) + Y .dist (f₁ .fun (x .rfun ε₂)) (f₂ .fun (x .rfun ε₂)) 104 | ≤⟨ +-mono-≤ (f₁ .non-expansive) (f₁≈f₂ .f≈f (x .rfun ε₂)) ⟩ 105 | X .dist (x .rfun ε₁) (x .rfun ε₂) + 0ℝ 106 | ≤⟨ +-mono-≤ (x .regular ε₁ ε₂) ≤-refl ⟩ 107 | rational+ (ε₁ ℚ⁺.+ ε₂) + 0ℝ 108 | ≈⟨ +-identityʳ (rational+ (ε₁ ℚ⁺.+ ε₂)) ⟩ 109 | rational+ (ε₁ ℚ⁺.+ ε₂) 110 | ∎ 111 | where open ≤-Reasoning 112 | 113 | open MetricSpace.category 114 | 115 | map-id : ∀ {X} → map {X} id ≈f id 116 | map-id {X} .f≈f x = 𝒞 X .refl {x} 117 | 118 | map-∘ : ∀ {X Y Z} {f : Y ⇒ Z} {g : X ⇒ Y} → 119 | map (f ∘ g) ≈f (map f ∘ map g) 120 | map-∘ {Z = Z}{f}{g} .f≈f a = 𝒞 Z .refl {map f .fun (map g .fun a)} 121 | 122 | ------------------------------------------------------------------------------ 123 | unit : ∀ {X} → X ⇒ 𝒞 X 124 | unit {X} .fun x .rfun _ = x 125 | unit {X} .fun x .regular ε₁ ε₂ = ≤-trans (X .refl) (0-least _) 126 | unit {X} .non-expansive {x}{y} = 127 | sup-least λ { (ε₁ , ε₂) → 128 | begin 129 | X .dist x y ⊖ (ε₁ ℚ⁺.+ ε₂) 130 | ≤⟨ ⊖-≤ ⟩ 131 | X .dist x y 132 | ∎ } 133 | where open ≤-Reasoning 134 | 135 | unit-natural : ∀ {X Y}{f : X ⇒ Y} → 136 | (unit ∘ f) ≈f (map f ∘ unit) 137 | unit-natural {X}{Y}{f} .f≈f x = (𝒞 Y) .refl {unit .fun (f .fun x)} 138 | 139 | ------------------------------------------------------------------------------ 140 | join : ∀ {X} → 𝒞 (𝒞 X) ⇒ 𝒞 X 141 | join {X} .fun x .rfun ε = x .rfun (ε /2) .rfun (ε /2) 142 | join {X} .fun x .regular ε₁ ε₂ = 143 | begin 144 | X .dist (x .rfun (ε₁ /2) .rfun (ε₁ /2)) (x .rfun (ε₂ /2) .rfun (ε₂ /2)) 145 | ≤⟨ ⊖-iso1 (≤-trans (sup-upper (ε₁ /2 , ε₂ /2)) (x .regular (ε₁ /2) (ε₂ /2))) ⟩ 146 | rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 147 | ≈⟨ rational⁺-+ (ε₁ /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 ℚ⁺.+ ε₂ /2) ⟩ 148 | rational+ ((ε₁ /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 149 | ≤⟨ rational-mono (ℚ⁺.fog-mono (ℚ⁺.≤-reflexive (ℚ⁺-interchange (ε₁ /2) (ε₂ /2) (ε₁ /2) (ε₂ /2)))) ⟩ 150 | rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 151 | ≤⟨ rational-mono (ℚ⁺.fog-mono (ℚ⁺.≤-reflexive (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂})))) ⟩ 152 | rational+ (ε₁ ℚ⁺.+ ε₂) 153 | ∎ 154 | where open ≤-Reasoning 155 | join {X} .non-expansive {a}{b} = 156 | sup-least (λ { (ε₁ , ε₂) → ⊖-iso2 157 | (begin 158 | X .dist (a .rfun (ε₁ /2) .rfun (ε₁ /2)) (b .rfun (ε₂ /2) .rfun (ε₂ /2)) 159 | ≤⟨ ⊖-eval ⟩ 160 | (X .dist (a .rfun (ε₁ /2) .rfun (ε₁ /2)) (b .rfun (ε₂ /2) .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 161 | ≤⟨ +-mono-≤ (⊖-mono ⊖-eval ℚ⁺.≤-refl) ≤-refl ⟩ 162 | (((X .dist (a .rfun (ε₁ /2) .rfun (ε₁ /2)) (b .rfun (ε₂ /2) .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 163 | ≤⟨ +-mono-≤ (⊖-+-out _ (rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) (ε₁ /2 ℚ⁺.+ ε₂ /2)) ≤-refl ⟩ 164 | (((X .dist (a .rfun (ε₁ /2) .rfun (ε₁ /2)) (b .rfun (ε₂ /2) .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 165 | ≈⟨ +-assoc _ (rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) (rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) ⟩ 166 | ((X .dist (a .rfun (ε₁ /2) .rfun (ε₁ /2)) (b .rfun (ε₂ /2) .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + (rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 167 | ≤⟨ +-mono-≤ (⊖-mono (sup-upper (ε₁ /2 , ε₂ /2)) ℚ⁺.≤-refl) (≤-reflexive (rational⁺-+ (ε₁ /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 ℚ⁺.+ ε₂ /2))) ⟩ 168 | (𝒞 X .dist (a .rfun (ε₁ /2)) (b .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ ((ε₁ /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 169 | ≈⟨ +-cong ≃-refl (rational-cong (ℚ⁺.fog-cong (ℚ⁺-interchange (ε₁ /2) (ε₂ /2) (ε₁ /2) (ε₂ /2)))) ⟩ 170 | (𝒞 X .dist (a .rfun (ε₁ /2)) (b .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 171 | ≈⟨ +-cong ≃-refl (rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂})))) ⟩ 172 | ((𝒞 X .dist (a .rfun (ε₁ /2)) (b .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ ℚ⁺.+ ε₂)) 173 | ≤⟨ +-mono-≤ (sup-upper (ε₁ /2 , ε₂ /2)) ≤-refl ⟩ 174 | dist (𝒞 (𝒞 X)) a b + rational+ (ε₁ ℚ⁺.+ ε₂) 175 | ∎) }) 176 | where open ≤-Reasoning 177 | 178 | join-natural : ∀ {X Y}{f : X ⇒ Y} → 179 | (join ∘ map (map f)) ≈f (map f ∘ join) 180 | join-natural {X}{Y}{f} .f≈f x = 𝒞 Y .refl {map f .fun (join .fun x)} 181 | 182 | join-unit : ∀ {X} → (join ∘ unit) ≈f id {𝒞 X} 183 | join-unit .f≈f a = 184 | sup-least (λ { (ε₁ , ε₂) → 185 | ≤-trans (⊖-mono (a .regular (ε₁ /2) ε₂) ℚ⁺.≤-refl) 186 | (⊖-iso2 (≤-trans (rational-mono (ℚ⁺.fog-mono (ℚ⁺.+-mono-≤ (ℚ⁺.half-≤ ε₁) (ℚ⁺.≤-refl {ε₂})))) 187 | (≤-reflexive (≃-sym (+-identityˡ (rational+ (ε₁ ℚ⁺.+ ε₂))))))) }) 188 | 189 | join-mapunit : ∀ {X} → (join ∘ map unit) ≈f id {𝒞 X} 190 | join-mapunit .f≈f a = join-unit .f≈f a 191 | 192 | join-join : ∀ {X} → (join ∘ map join) ≈f (join ∘ join {𝒞 X}) 193 | join-join {X} .f≈f x = 194 | 𝒞-≈ {X} {join .fun (map join .fun x)} {join .fun (join .fun x)} λ ε₁ ε₂ → 195 | ≤-trans (⊖-iso1 (≤-trans (sup-upper (ε₁ /2 /2 , ε₂ /2)) (⊖-iso1 (≤-trans (sup-upper (ε₁ /2 /2 , ε₂ /2 /2)) (x .regular (ε₁ /2) (ε₂ /2 /2)))))) (eq ε₁ ε₂) 196 | where 197 | open ≤-Reasoning 198 | open import Algebra.Solver.CommutativeSemigroup (ℚ⁺.+-commutativeSemigroup) 199 | a = v# 0; b = v# 1; c = v# 2; d = v# 3 200 | eq : ∀ ε₁ ε₂ → ((rational+ (ε₁ /2 ℚ⁺.+ (ε₂ /2) /2) + rational+ ((ε₁ /2) /2 ℚ⁺.+ (ε₂ /2) /2)) + rational+ ((ε₁ /2) /2 ℚ⁺.+ ε₂ /2)) ≤ rational+ (ε₁ ℚ⁺.+ ε₂) 201 | eq ε₁ ε₂ = 202 | begin 203 | (rational+ (ε₁ /2 ℚ⁺.+ (ε₂ /2) /2) + rational+ ((ε₁ /2) /2 ℚ⁺.+ (ε₂ /2) /2)) + rational+ ((ε₁ /2) /2 ℚ⁺.+ ε₂ /2) 204 | ≈⟨ +-cong (rational⁺-+ (ε₁ /2 ℚ⁺.+ (ε₂ /2) /2) ((ε₁ /2) /2 ℚ⁺.+ (ε₂ /2) /2)) ≃-refl ⟩ 205 | rational+ ((ε₁ /2 ℚ⁺.+ (ε₂ /2) /2) ℚ⁺.+ ((ε₁ /2) /2 ℚ⁺.+ (ε₂ /2) /2)) + rational+ ((ε₁ /2) /2 ℚ⁺.+ ε₂ /2) 206 | ≈⟨ rational⁺-+ ((ε₁ /2 ℚ⁺.+ (ε₂ /2) /2) ℚ⁺.+ ((ε₁ /2) /2 ℚ⁺.+ (ε₂ /2) /2)) ((ε₁ /2) /2 ℚ⁺.+ ε₂ /2) ⟩ 207 | rational+ ((ε₁ /2 ℚ⁺.+ (ε₂ /2) /2) ℚ⁺.+ ((ε₁ /2) /2 ℚ⁺.+ (ε₂ /2) /2) ℚ⁺.+ ((ε₁ /2) /2 ℚ⁺.+ ε₂ /2)) 208 | ≈⟨ rational-cong (ℚ⁺.fog-cong (prove 4 (((a ⊕ b) ⊕ (c ⊕ b)) ⊕ (c ⊕ d)) ((a ⊕ (c ⊕ c)) ⊕ (d ⊕ (b ⊕ b))) (ε₁ /2 ∷ ε₂ /2 /2 ∷ ε₁ /2 /2 ∷ ε₂ /2 ∷ []))) ⟩ 209 | rational+ ((ε₁ /2 ℚ⁺.+ (ε₁ /2 /2 ℚ⁺.+ ε₁ /2 /2)) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ (ε₂ /2 /2 ℚ⁺.+ ε₂ /2 /2))) 210 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.+-congʳ (ε₁ /2) (ℚ⁺.half+half {ε₁ /2})) (ℚ⁺.+-congʳ (ε₂ /2) (ℚ⁺.half+half {ε₂ /2})))) ⟩ 211 | rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 212 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂}))) ⟩ 213 | rational+ (ε₁ ℚ⁺.+ ε₂) 214 | -- ≈⟨ ≃-sym (+-identityˡ (rational+ (ε₁ ℚ⁺.+ ε₂))) ⟩ 215 | -- 0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂) 216 | ∎ 217 | 218 | ------------------------------------------------------------------------------ 219 | -- Idempotency 220 | -- 221 | -- This shows that join : 𝒞 (𝒞 X) ≅ 𝒞 X : unit 222 | 223 | unit-join : ∀ {X} → (unit ∘ join) ≈f id {𝒞 (𝒞 X)} 224 | unit-join .f≈f x = 225 | sup-least λ { (ε₁ , ε₂) → 226 | ⊖-iso2 (sup-least λ { (ε₁' , ε₂') → 227 | ⊖-iso2 (≤-trans (⊖-iso1 (≤-trans (sup-upper (ε₁' /2 , ε₂')) (x .regular (ε₁' /2) ε₂))) (ineq ε₁ ε₂ ε₁' ε₂')) }) } 228 | where 229 | open ≤-Reasoning 230 | open import Algebra.Solver.CommutativeSemigroup (ℚ⁺.+-commutativeSemigroup) 231 | a = v# 0; b = v# 1; c = v# 2; d = v# 3 232 | ineq : ∀ ε₁ ε₂ ε₁' ε₂' → (rational+ (ε₁' /2 ℚ⁺.+ ε₂) + rational+ (ε₁' /2 ℚ⁺.+ ε₂')) ≤ ((0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂)) + rational+ (ε₁' ℚ⁺.+ ε₂')) 233 | ineq ε₁ ε₂ ε₁' ε₂' = 234 | begin 235 | rational+ (ε₁' /2 ℚ⁺.+ ε₂) + rational+ (ε₁' /2 ℚ⁺.+ ε₂') 236 | ≈⟨ rational⁺-+ (ε₁' /2 ℚ⁺.+ ε₂) (ε₁' /2 ℚ⁺.+ ε₂') ⟩ 237 | rational+ ((ε₁' /2 ℚ⁺.+ ε₂) ℚ⁺.+ (ε₁' /2 ℚ⁺.+ ε₂')) 238 | ≈⟨ rational-cong (ℚ⁺.fog-cong (prove 3 ((a ⊕ b) ⊕ (a ⊕ c)) (b ⊕ ((a ⊕ a) ⊕ c)) (ε₁' /2 ∷ ε₂ ∷ ε₂' ∷ []))) ⟩ 239 | rational+ (ε₂ ℚ⁺.+ ((ε₁' /2 ℚ⁺.+ ε₁' /2) ℚ⁺.+ ε₂')) 240 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-congʳ ε₂ (ℚ⁺.+-congˡ ε₂' (ℚ⁺.half+half {ε₁'})))) ⟩ 241 | rational+ (ε₂ ℚ⁺.+ (ε₁' ℚ⁺.+ ε₂')) 242 | ≤⟨ rational-mono (ℚ⁺.fog-mono (ℚ⁺.+-mono-≤ (ℚ⁺.+-increasing {ε₂} {ε₁}) (ℚ⁺.≤-refl {ε₁' ℚ⁺.+ ε₂'}))) ⟩ 243 | rational+ ((ε₂ ℚ⁺.+ ε₁) ℚ⁺.+ (ε₁' ℚ⁺.+ ε₂')) 244 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-congˡ (ε₁' ℚ⁺.+ ε₂') (ℚ⁺.+-comm ε₂ ε₁))) ⟩ 245 | rational+ ((ε₁ ℚ⁺.+ ε₂) ℚ⁺.+ (ε₁' ℚ⁺.+ ε₂')) 246 | ≈⟨ ≃-sym (rational⁺-+ (ε₁ ℚ⁺.+ ε₂) (ε₁' ℚ⁺.+ ε₂')) ⟩ 247 | rational+ (ε₁ ℚ⁺.+ ε₂) + rational+ (ε₁' ℚ⁺.+ ε₂') 248 | ≈⟨ +-cong (≃-sym (+-identityˡ (rational+ (ε₁ ℚ⁺.+ ε₂)))) ≃-refl ⟩ 249 | (0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂)) + rational+ (ε₁' ℚ⁺.+ ε₂') 250 | ∎ 251 | 252 | ------------------------------------------------------------------------------ 253 | -- This is a monoidal monad, with respect to the monoidal product 254 | 255 | -- FIXME: is this true for any norm-product? 256 | 257 | open import MetricSpace.MonoidalProduct 258 | open import MetricSpace.Terminal 259 | 260 | monoidal-⊗ : ∀ {X Y} → (𝒞 X ⊗ 𝒞 Y) ⇒ 𝒞 (X ⊗ Y) 261 | monoidal-⊗ .fun (x , y) .rfun ε = x .rfun (ε /2) , y .rfun (ε /2) 262 | monoidal-⊗ {X}{Y} .fun (x , y) .regular ε₁ ε₂ = 263 | begin 264 | (X .dist (x .rfun (ε₁ /2)) (x .rfun (ε₂ /2)) + Y .dist (y .rfun (ε₁ /2)) (y .rfun (ε₂ /2))) 265 | ≤⟨ +-mono-≤ (x .regular (ε₁ /2) (ε₂ /2)) (y .regular (ε₁ /2) (ε₂ /2)) ⟩ 266 | rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 267 | ≈⟨ rational⁺-+ (ε₁ /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 ℚ⁺.+ ε₂ /2) ⟩ 268 | rational+ ((ε₁ /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 269 | ≤⟨ rational-mono (ℚ⁺.fog-mono (ℚ⁺.≤-reflexive (ℚ⁺-interchange (ε₁ /2) (ε₂ /2) (ε₁ /2) (ε₂ /2)))) ⟩ 270 | rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 271 | ≤⟨ rational-mono (ℚ⁺.fog-mono (ℚ⁺.≤-reflexive (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂})))) ⟩ 272 | rational+ (ε₁ ℚ⁺.+ ε₂) 273 | ∎ 274 | where open ≤-Reasoning 275 | monoidal-⊗ {X}{Y} .non-expansive {x₁ , y₁} {x₂ , y₂} = 276 | sup-least λ { (ε₁ , ε₂) → ⊖-iso2 277 | (begin 278 | X .dist (x₁ .rfun (ε₁ /2)) (x₂ .rfun (ε₂ /2)) + Y .dist (y₁ .rfun (ε₁ /2)) (y₂ .rfun (ε₂ /2)) 279 | ≤⟨ +-mono-≤ ⊖-eval ⊖-eval ⟩ 280 | ((X .dist (x₁ .rfun (ε₁ /2)) (x₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 281 | + ((Y .dist (y₁ .rfun (ε₁ /2)) (y₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 282 | ≈⟨ interchange (X .dist (x₁ .rfun (ε₁ /2)) (x₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) (rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 283 | (Y .dist (y₁ .rfun (ε₁ /2)) (y₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) (rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) ⟩ 284 | ( (X .dist (x₁ .rfun (ε₁ /2)) (x₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 285 | + (Y .dist (y₁ .rfun (ε₁ /2)) (y₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2))) 286 | + (rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 287 | ≈⟨ +-cong ≃-refl (rational⁺-+ (ε₁ /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 ℚ⁺.+ ε₂ /2)) ⟩ 288 | ( (X .dist (x₁ .rfun (ε₁ /2)) (x₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 289 | + (Y .dist (y₁ .rfun (ε₁ /2)) (y₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2))) 290 | + (rational+ ((ε₁ /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2))) 291 | ≤⟨ +-mono-≤ ≤-refl (rational-mono (ℚ⁺.fog-mono (ℚ⁺.≤-reflexive (ℚ⁺-interchange (ε₁ /2) (ε₂ /2) (ε₁ /2) (ε₂ /2))))) ⟩ 292 | ( (X .dist (x₁ .rfun (ε₁ /2)) (x₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 293 | + (Y .dist (y₁ .rfun (ε₁ /2)) (y₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2))) 294 | + (rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2))) 295 | ≤⟨ +-mono-≤ ≤-refl (rational-mono (ℚ⁺.fog-mono (ℚ⁺.≤-reflexive (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂}))))) ⟩ 296 | ( (X .dist (x₁ .rfun (ε₁ /2)) (x₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 297 | + (Y .dist (y₁ .rfun (ε₁ /2)) (y₂ .rfun (ε₂ /2)) ⊖ (ε₁ /2 ℚ⁺.+ ε₂ /2))) + rational+ (ε₁ ℚ⁺.+ ε₂) 298 | ≤⟨ +-mono-≤ (+-mono-≤ (sup-upper (ε₁ /2 , ε₂ /2)) (sup-upper (ε₁ /2 , ε₂ /2))) ≤-refl ⟩ 299 | (reg-dist x₁ x₂ + reg-dist y₁ y₂) + rational+ (ε₁ ℚ⁺.+ ε₂) 300 | ∎) } 301 | where open ≤-Reasoning 302 | 303 | monoidal-natural : ∀ {X X' Y Y'} → (f : X ⇒ X') (g : Y ⇒ Y') → 304 | (monoidal-⊗ ∘ (map f ⊗f map g)) ≈f (map (f ⊗f g) ∘ monoidal-⊗) 305 | monoidal-natural {X}{X'}{Y}{Y'} f g .f≈f (x , y) = 306 | sup-least λ { (ε₁ , ε₂) → ⊖-iso2 307 | (begin 308 | X' .dist (f .fun (x .rfun (ε₁ /2))) (f .fun (x .rfun (ε₂ /2))) + Y' .dist (g .fun (y .rfun (ε₁ /2))) (g .fun (y .rfun (ε₂ /2))) 309 | ≤⟨ +-mono-≤ (f .non-expansive) (g .non-expansive) ⟩ 310 | (X .dist (x .rfun (ε₁ /2)) (x .rfun (ε₂ /2)) + Y .dist (y .rfun (ε₁ /2)) (y .rfun (ε₂ /2))) 311 | ≤⟨ +-mono-≤ (x .regular (ε₁ /2) (ε₂ /2)) (y .regular (ε₁ /2) (ε₂ /2)) ⟩ 312 | rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 313 | ≈⟨ rational⁺-+ (ε₁ /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 ℚ⁺.+ ε₂ /2) ⟩ 314 | rational+ ((ε₁ /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 315 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺-interchange (ε₁ /2) (ε₂ /2) (ε₁ /2) (ε₂ /2))) ⟩ 316 | rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 317 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂}))) ⟩ 318 | rational+ (ε₁ ℚ⁺.+ ε₂) 319 | ≈⟨ ≃-sym (+-identityˡ (rational+ (ε₁ ℚ⁺.+ ε₂))) ⟩ 320 | 0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂) 321 | ∎) } 322 | where open ≤-Reasoning 323 | 324 | monoidal-assoc : ∀ {X Y Z} → 325 | (monoidal-⊗ ∘ (monoidal-⊗ ⊗f id) ∘ ⊗-assoc {𝒞 X} {𝒞 Y} {𝒞 Z}) ≈f (map ⊗-assoc ∘ monoidal-⊗ ∘ (id ⊗f monoidal-⊗)) 326 | monoidal-assoc {X}{Y}{Z} .f≈f (x , (y , z)) = 327 | sup-least λ { (ε₁ , ε₂) → ⊖-iso2 328 | (begin 329 | (X .dist (x .rfun ((ε₁ /2) /2)) (x .rfun (ε₂ /2)) + Y .dist (y .rfun ((ε₁ /2) /2)) (y .rfun ((ε₂ /2) /2))) + Z .dist (z .rfun (ε₁ /2)) (z .rfun ((ε₂ /2) /2)) 330 | ≤⟨ +-mono-≤ (+-mono-≤ (x .regular (ε₁ /2 /2) (ε₂ /2)) (y .regular (ε₁ /2 /2) (ε₂ /2 /2))) (z .regular (ε₁ /2) (ε₂ /2 /2)) ⟩ 331 | (rational+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2 /2) 332 | ≈⟨ +-cong (rational⁺-+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) ≃-refl ⟩ 333 | rational+ ((ε₁ /2 /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2 /2) 334 | ≈⟨ rational⁺-+ ((ε₁ /2 /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) (ε₁ /2 ℚ⁺.+ ε₂ /2 /2) ⟩ 335 | rational+ (((ε₁ /2 /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2 /2)) 336 | ≈⟨ rational-cong (ℚ⁺.fog-cong 337 | (prove 4 (((b ⊕ c) ⊕ (b ⊕ d)) ⊕ (a ⊕ d)) ((a ⊕ (b ⊕ b)) ⊕ (c ⊕ (d ⊕ d))) (ε₁ /2 ∷ ε₁ /2 /2 ∷ ε₂ /2 ∷ ε₂ /2 /2 ∷ []))) ⟩ 338 | rational+ ((ε₁ /2 ℚ⁺.+ (ε₁ /2 /2 ℚ⁺.+ ε₁ /2 /2)) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ (ε₂ /2 /2 ℚ⁺.+ ε₂ /2 /2))) 339 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.+-congʳ (ε₁ /2) (ℚ⁺.half+half {ε₁ /2})) (ℚ⁺.+-congʳ (ε₂ /2) (ℚ⁺.half+half {ε₂ /2})))) ⟩ 340 | rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 341 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂}))) ⟩ 342 | rational+ (ε₁ ℚ⁺.+ ε₂) 343 | ≈⟨ ≃-sym (+-identityˡ (rational+ (ε₁ ℚ⁺.+ ε₂))) ⟩ 344 | 0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂) 345 | ∎) } 346 | where open ≤-Reasoning 347 | open import Algebra.Solver.CommutativeSemigroup (ℚ⁺.+-commutativeSemigroup) 348 | a = v# 0; b = v# 1; c = v# 2; d = v# 3 349 | 350 | monoidal-symmetry : ∀ {X Y} → 351 | (monoidal-⊗ {Y} {X} ∘ ⊗-symmetry) ≈f (map ⊗-symmetry ∘ monoidal-⊗) 352 | monoidal-symmetry {X}{Y} .f≈f (x , y) = 353 | sup-least λ { (ε₁ , ε₂) → ⊖-iso2 354 | (begin 355 | Y .dist (y .rfun (ε₁ /2)) (y .rfun (ε₂ /2)) + X .dist (x .rfun (ε₁ /2)) (x .rfun (ε₂ /2)) 356 | ≤⟨ +-mono-≤ (y .regular (ε₁ /2) (ε₂ /2)) (x .regular (ε₁ /2) (ε₂ /2)) ⟩ 357 | rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 358 | ≈⟨ rational⁺-+ (ε₁ /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 ℚ⁺.+ ε₂ /2) ⟩ 359 | rational+ ((ε₁ /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 360 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺-interchange (ε₁ /2) (ε₂ /2) (ε₁ /2) (ε₂ /2))) ⟩ 361 | rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 362 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂}))) ⟩ 363 | rational+ (ε₁ ℚ⁺.+ ε₂) 364 | ≈⟨ ≃-sym (+-identityˡ (rational+ (ε₁ ℚ⁺.+ ε₂))) ⟩ 365 | 0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂) 366 | ∎) 367 | } 368 | where open ≤-Reasoning 369 | 370 | monoidal-left-unit : ∀ {X} → (monoidal-⊗ {⊤ₘ}{X} ∘ (unit ⊗f id)) ≈f (map left-unit⁻¹ ∘ left-unit) 371 | monoidal-left-unit {X} .f≈f (tt , x) = 372 | sup-least λ { (ε₁ , ε₂) → ⊖-iso2 373 | (begin 374 | (0ℝ + X .dist (x .rfun (ε₁ /2)) (x .rfun ε₂)) 375 | ≤⟨ +-mono-≤ ≤-refl (x .regular (ε₁ /2) ε₂) ⟩ 376 | (0ℝ + rational+ (ε₁ /2 ℚ⁺.+ ε₂)) 377 | ≤⟨ +-mono-≤ ≤-refl (rational-mono (ℚ⁺.fog-mono (ℚ⁺.+-mono-≤ (ℚ⁺.half-≤ ε₁) (ℚ⁺.≤-refl {ε₂})))) ⟩ 378 | (0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂)) 379 | ∎) } 380 | where open ≤-Reasoning 381 | 382 | monoidal-unit : ∀ {X Y} → (monoidal-⊗ {X} {Y} ∘ (unit ⊗f unit)) ≈f unit 383 | monoidal-unit {X}{Y} .f≈f (x , y) = 384 | sup-least (λ { (ε₁ , ε₂) → ⊖-iso2 385 | (begin 386 | X .dist x x + Y .dist y y 387 | ≤⟨ +-mono-≤ (X .refl) (Y .refl) ⟩ 388 | 0ℝ + 0ℝ 389 | ≤⟨ +-mono-≤ ≤-refl (0-least (rational+ (ε₁ ℚ⁺.+ ε₂))) ⟩ 390 | 0ℝ + rational+ (ε₁ ℚ⁺.+ ε₂) 391 | ∎) }) 392 | where open ≤-Reasoning 393 | 394 | monoidal-join : ∀ {X Y} → (monoidal-⊗ {X}{Y} ∘ (join ⊗f join)) ≈f (join ∘ map monoidal-⊗ ∘ monoidal-⊗) 395 | monoidal-join {X}{Y} .f≈f (x , y) = 396 | 𝒞-≈ {X ⊗ Y} {monoidal-⊗ .fun (join .fun x , join .fun y)} 397 | {join .fun (map monoidal-⊗ .fun (monoidal-⊗ .fun (x , y)))} 398 | λ ε₁ ε₂ → 399 | begin 400 | X .dist (x .rfun ((ε₁ /2) /2) .rfun ((ε₁ /2) /2)) (x .rfun ((ε₂ /2) /2) .rfun ((ε₂ /2) /2)) 401 | + Y .dist (y .rfun ((ε₁ /2) /2) .rfun ((ε₁ /2) /2)) (y .rfun ((ε₂ /2) /2) .rfun ((ε₂ /2) /2)) 402 | ≤⟨ +-mono-≤ (⊖-iso1 (≤-trans (sup-upper (ε₁ /2 /2 , ε₂ /2 /2)) (x .regular (ε₁ /2 /2) (ε₂ /2 /2)))) 403 | (⊖-iso1 (≤-trans (sup-upper (ε₁ /2 /2 , ε₂ /2 /2)) (y .regular (ε₁ /2 /2) (ε₂ /2 /2)))) ⟩ 404 | (rational+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2) + rational+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) + (rational+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2) + rational+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) 405 | ≈⟨ +-cong (rational⁺-+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2) (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) (rational⁺-+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2) (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) ⟩ 406 | rational+ ((ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2) ℚ⁺.+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) + rational+ ((ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2) ℚ⁺.+ (ε₁ /2 /2 ℚ⁺.+ ε₂ /2 /2)) 407 | ≈⟨ +-cong (rational-cong (ℚ⁺.fog-cong (ℚ⁺-interchange (ε₁ /2 /2) (ε₂ /2 /2) (ε₁ /2 /2) (ε₂ /2 /2)))) 408 | (rational-cong (ℚ⁺.fog-cong (ℚ⁺-interchange (ε₁ /2 /2) (ε₂ /2 /2) (ε₁ /2 /2) (ε₂ /2 /2)))) ⟩ 409 | rational+ ((ε₁ /2 /2 ℚ⁺.+ ε₁ /2 /2) ℚ⁺.+ (ε₂ /2 /2 ℚ⁺.+ ε₂ /2 /2)) + rational+ ((ε₁ /2 /2 ℚ⁺.+ ε₁ /2 /2) ℚ⁺.+ (ε₂ /2 /2 ℚ⁺.+ ε₂ /2 /2)) 410 | ≈⟨ +-cong (rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁ /2}) (ℚ⁺.half+half {ε₂ /2})))) 411 | (rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁ /2}) (ℚ⁺.half+half {ε₂ /2})))) ⟩ 412 | rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) + rational+ (ε₁ /2 ℚ⁺.+ ε₂ /2) 413 | ≈⟨ rational⁺-+ (ε₁ /2 ℚ⁺.+ ε₂ /2) (ε₁ /2 ℚ⁺.+ ε₂ /2) ⟩ 414 | rational+ ((ε₁ /2 ℚ⁺.+ ε₂ /2) ℚ⁺.+ (ε₁ /2 ℚ⁺.+ ε₂ /2)) 415 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺-interchange (ε₁ /2) (ε₂ /2) (ε₁ /2) (ε₂ /2))) ⟩ 416 | rational+ ((ε₁ /2 ℚ⁺.+ ε₁ /2) ℚ⁺.+ (ε₂ /2 ℚ⁺.+ ε₂ /2)) 417 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.+-cong (ℚ⁺.half+half {ε₁}) (ℚ⁺.half+half {ε₂}))) ⟩ 418 | rational+ (ε₁ ℚ⁺.+ ε₂) 419 | ∎ 420 | where open ≤-Reasoning 421 | 422 | {- 423 | monoidal-⊗⁻¹ : ∀ {X Y} → 𝒞 (X ⊗ Y) ⇒ (𝒞 X ⊗ 𝒞 Y) 424 | monoidal-⊗⁻¹ .fun x .proj₁ .rfun ε = x .rfun ε .proj₁ 425 | monoidal-⊗⁻¹ .fun x .proj₁ .regular ε₁ ε₂ = {!!} 426 | monoidal-⊗⁻¹ .fun x .proj₂ .rfun ε = x .rfun ε .proj₂ 427 | monoidal-⊗⁻¹ .fun x .proj₂ .regular ε₁ ε₂ = {!!} 428 | monoidal-⊗⁻¹ .non-expansive = {!!} 429 | -} 430 | 431 | ------------------------------------------------------------------------------ 432 | open import MetricSpace.Scaling 433 | 434 | distr : ∀ {X} q → ![ q ] (𝒞 X) ⇒ 𝒞 (![ q ] X) 435 | distr {X} q .fun x .rfun ε = x .rfun (1/ q ℚ⁺.* ε) 436 | distr {X} q .fun x .regular ε₁ ε₂ = 437 | begin 438 | rational+ q * X .dist (x .rfun (1/ q ℚ⁺.* ε₁)) (x .rfun (1/ q ℚ⁺.* ε₂)) 439 | ≤⟨ *-mono-≤ ≤-refl (x .regular (1/ q ℚ⁺.* ε₁) (1/ q ℚ⁺.* ε₂)) ⟩ 440 | rational+ q * rational+ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂)) 441 | ≈⟨ *-cong ≃-refl (rational-cong (ℚ⁺.fog-cong (ℚ⁺.≃-sym (ℚ⁺.*-distribˡ-+ (1/ q) ε₁ ε₂)))) ⟩ 442 | rational+ q * rational+ (1/ q ℚ⁺.* (ε₁ ℚ⁺.+ ε₂)) 443 | ≈⟨ rational⁺-* q (1/ q ℚ⁺.* (ε₁ ℚ⁺.+ ε₂)) ⟩ 444 | rational+ (q ℚ⁺.* (1/ q ℚ⁺.* (ε₁ ℚ⁺.+ ε₂))) 445 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.≃-sym (ℚ⁺.*-assoc q (1/ q) (ε₁ ℚ⁺.+ ε₂)))) ⟩ 446 | rational+ ((q ℚ⁺.* 1/ q) ℚ⁺.* (ε₁ ℚ⁺.+ ε₂)) 447 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.*-cong (ℚ⁺.*-inverseʳ q) (ℚ⁺.≃-refl {ε₁ ℚ⁺.+ ε₂}))) ⟩ 448 | rational+ (1ℚ⁺ ℚ⁺.* (ε₁ ℚ⁺.+ ε₂)) 449 | ≈⟨ rational-cong (ℚ⁺.fog-cong (ℚ⁺.*-identityˡ (ε₁ ℚ⁺.+ ε₂))) ⟩ 450 | rational+ (ε₁ ℚ⁺.+ ε₂) 451 | ∎ 452 | where open ≤-Reasoning 453 | distr {X} q .non-expansive {x} {y} = 454 | sup-least λ { (ε₁ , ε₂) → ⊖-iso2 455 | (begin 456 | rational+ q * X .dist (x .rfun (1/ q ℚ⁺.* ε₁)) (y .rfun (1/ q ℚ⁺.* ε₂)) 457 | ≤⟨ *-mono-≤ ≤-refl ⊖-eval ⟩ 458 | rational+ q * ((X .dist (x .rfun (1/ q ℚ⁺.* ε₁)) (y .rfun (1/ q ℚ⁺.* ε₂)) ⊖ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) + rational+ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) 459 | ≈⟨ *-distribˡ-+ (rational+ q) 460 | (X .dist (x .rfun (1/ q ℚ⁺.* ε₁)) (y .rfun (1/ q ℚ⁺.* ε₂)) ⊖ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) 461 | (rational+ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) ⟩ 462 | (rational+ q * (X .dist (x .rfun (1/ q ℚ⁺.* ε₁)) (y .rfun (1/ q ℚ⁺.* ε₂)) ⊖ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂)))) + 463 | (rational+ q * rational+ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) 464 | ≤⟨ +-mono-≤ (*-mono-≤ ≤-refl (sup-upper (1/ q ℚ⁺.* ε₁ , 1/ q ℚ⁺.* ε₂))) ≤-refl ⟩ 465 | dist (![ q ] (𝒞 X)) x y + (rational+ q * rational+ ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) 466 | ≈⟨ +-cong ≃-refl (rational⁺-* q ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) ⟩ 467 | dist (![ q ] (𝒞 X)) x y + rational+ (q ℚ⁺.* ((1/ q ℚ⁺.* ε₁) ℚ⁺.+ (1/ q ℚ⁺.* ε₂))) 468 | ≈⟨ +-cong ≃-refl (rational-cong (ℚ⁺.fog-cong (ℚ⁺.*-cong (ℚ⁺.≃-refl {q}) (ℚ⁺.≃-sym (ℚ⁺.*-distribˡ-+ (1/ q) ε₁ ε₂))))) ⟩ 469 | dist (![ q ] (𝒞 X)) x y + rational+ (q ℚ⁺.* (1/ q ℚ⁺.* (ε₁ ℚ⁺.+ ε₂))) 470 | ≈⟨ +-cong ≃-refl (rational-cong (ℚ⁺.fog-cong (ℚ⁺.≃-sym (ℚ⁺.*-assoc q (1/ q) (ε₁ ℚ⁺.+ ε₂))))) ⟩ 471 | dist (![ q ] (𝒞 X)) x y + rational+ ((q ℚ⁺.* 1/ q) ℚ⁺.* (ε₁ ℚ⁺.+ ε₂)) 472 | ≈⟨ +-cong ≃-refl (rational-cong (ℚ⁺.fog-cong (ℚ⁺.*-cong (ℚ⁺.*-inverseʳ q) (ℚ⁺.≃-refl {ε₁ ℚ⁺.+ ε₂})))) ⟩ 473 | dist (![ q ] (𝒞 X)) x y + rational+ (1ℚ⁺ ℚ⁺.* (ε₁ ℚ⁺.+ ε₂)) 474 | ≈⟨ +-cong ≃-refl (rational-cong (ℚ⁺.fog-cong (ℚ⁺.*-identityˡ (ε₁ ℚ⁺.+ ε₂)))) ⟩ 475 | dist (![ q ] (𝒞 X)) x y + rational+ (ε₁ ℚ⁺.+ ε₂) 476 | ∎) } 477 | where open ≤-Reasoning 478 | -------------------------------------------------------------------------------- /src/MetricSpace/ConvexAlgebras.agda: -------------------------------------------------------------------------------- 1 | module MetricSpace.ConvexAlgebras where 2 | 3 | open import Data.Product using (_,_; Σ-syntax; proj₁; proj₂; _×_) 4 | open import Data.Unit using (tt) 5 | import Data.Real.UpperClosed as ℝᵘ 6 | import Data.Integer as ℤ 7 | import Data.Nat as ℕ 8 | open import Data.Rational.Unnormalised as ℚ using () renaming (ℚᵘ to ℚ; 0ℚᵘ to 0ℚ; 1ℚᵘ to 1ℚ) 9 | import Data.Rational.Unnormalised.Properties as ℚ 10 | open import MetricSpace 11 | 12 | open import Data.Rational.Unnormalised.Positive as ℚ⁺ using (ℚ⁺; 1ℚ⁺) 13 | 14 | open MSpc 15 | 16 | record ℚ⟨0,1⟩ : Set where 17 | field 18 | num : ℚ 19 | 0 MSpc 47 | -------------------------------------------------------------------------------- /src/MetricSpace/InternalHom.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module MetricSpace.InternalHom where 4 | 5 | open import Data.Product using (_,_) 6 | open import MetricSpace 7 | open import MetricSpace.MonoidalProduct 8 | open import Data.Real.UpperClosed 9 | 10 | open MSpc 11 | open _⇒_ 12 | 13 | _⊸_ : MSpc → MSpc → MSpc 14 | (X ⊸ Y) .Carrier = X ⇒ Y 15 | (X ⊸ Y) .dist f g = sup (X .Carrier) λ x → Y .dist (f .fun x) (g .fun x) 16 | (X ⊸ Y) .refl = sup-least (λ x → Y .refl) 17 | (X ⊸ Y) .sym = sup-mono-≤ (λ x → x) λ x → Y .sym 18 | (X ⊸ Y) .triangle {f}{g}{h} = 19 | sup-least λ x → ≤-trans (Y .triangle) (+-mono-≤ (sup-upper x) (sup-upper x)) 20 | 21 | Λ : ∀ {X Y Z} → (X ⊗ Y) ⇒ Z → X ⇒ (Y ⊸ Z) 22 | Λ f .fun x .fun y = f .fun (x , y) 23 | Λ {X}{Y}{Z} f .fun x .non-expansive {y₁}{y₂} = 24 | begin 25 | Z .dist (f .fun (x , y₁)) (f .fun (x , y₂)) 26 | ≤⟨ f .non-expansive ⟩ 27 | (X ⊗ Y) .dist (x , y₁) (x , y₂) 28 | ≤⟨ +-mono-≤ (X .refl) ≤-refl ⟩ 29 | 0ℝ + Y .dist y₁ y₂ 30 | ≈⟨ +-identityˡ (Y .dist y₁ y₂) ⟩ 31 | Y .dist y₁ y₂ 32 | ∎ 33 | where open ≤-Reasoning 34 | Λ {X}{Y}{Z} f .non-expansive {x₁}{x₂}= 35 | sup-least (λ y → begin 36 | Z .dist (f .fun (x₁ , y)) (f .fun (x₂ , y)) 37 | ≤⟨ f .non-expansive ⟩ 38 | (X ⊗ Y) .dist (x₁ , y) (x₂ , y) 39 | ≤⟨ +-mono-≤ ≤-refl (Y .refl) ⟩ 40 | X .dist x₁ x₂ + 0ℝ 41 | ≈⟨ +-identityʳ (X .dist x₁ x₂) ⟩ 42 | X .dist x₁ x₂ 43 | ∎) 44 | where open ≤-Reasoning 45 | 46 | Λ⁻¹ : ∀ {X Y Z} → X ⇒ (Y ⊸ Z) → (X ⊗ Y) ⇒ Z 47 | Λ⁻¹ f .fun (x , y) = f .fun x .fun y 48 | Λ⁻¹ {X}{Y}{Z} f .non-expansive {x₁ , y₁}{x₂ , y₂} = 49 | begin 50 | Z .dist (f .fun x₁ .fun y₁) (f .fun x₂ .fun y₂) 51 | ≤⟨ Z .triangle ⟩ 52 | Z .dist (f .fun x₁ .fun y₁) (f .fun x₂ .fun y₁) + Z .dist (f .fun x₂ .fun y₁) (f .fun x₂ .fun y₂) 53 | ≤⟨ +-mono-≤ (sup-upper y₁) ≤-refl ⟩ 54 | (Y ⊸ Z) .dist (f .fun x₁) (f .fun x₂) + Z .dist (f .fun x₂ .fun y₁) (f .fun x₂ .fun y₂) 55 | ≤⟨ +-mono-≤ (f .non-expansive) (f .fun x₂ .non-expansive) ⟩ 56 | (X ⊗ Y) .dist (x₁ , y₁) (x₂ , y₂) 57 | ∎ 58 | where open ≤-Reasoning 59 | 60 | open MetricSpace.category 61 | 62 | eval : ∀ {X Y} → ((X ⊸ Y) ⊗ X) ⇒ Y 63 | eval = Λ⁻¹ id 64 | 65 | 66 | 67 | -- FIXME: 68 | -- 1. _⊸_ is a bifunuctor 69 | -- 2. Λ and Λ⁻¹ are natural 70 | -- 3. Λ and Λ⁻¹ are mutually inverse 71 | -------------------------------------------------------------------------------- /src/MetricSpace/MonoidalProduct.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module MetricSpace.MonoidalProduct where 4 | 5 | open import Data.Product using (_×_; _,_) 6 | open import Data.Unit using (tt) 7 | open import MetricSpace 8 | open import MetricSpace.Category 9 | open import MetricSpace.Terminal 10 | open import MetricSpace.CartesianProduct 11 | open import Data.Real.UpperClosed 12 | 13 | open MSpc 14 | open _⇒_ 15 | open _≈f_ 16 | 17 | -- FIXME: parameterise this by the norm being used. Construction of 18 | -- the monoidal structure ought to be generic. 19 | 20 | _⊗_ : MSpc → MSpc → MSpc 21 | (X ⊗ Y) .Carrier = 22 | X .Carrier × Y .Carrier 23 | (X ⊗ Y) .dist (x₁ , y₁) (x₂ , y₂) = 24 | X .dist x₁ x₂ + Y .dist y₁ y₂ 25 | (X ⊗ Y) .refl {x , y} = 26 | begin 27 | X .dist x x + Y .dist y y 28 | ≤⟨ +-mono-≤ (X .refl) (Y .refl) ⟩ 29 | 0ℝ + 0ℝ 30 | ≈⟨ +-identityˡ 0ℝ ⟩ 31 | 0ℝ 32 | ∎ 33 | where open ≤-Reasoning 34 | (X ⊗ Y) .sym {x₁ , y₁} {x₂ , y₂} = 35 | begin 36 | X .dist x₁ x₂ + Y .dist y₁ y₂ 37 | ≤⟨ +-mono-≤ (X .sym) (Y .sym) ⟩ 38 | X .dist x₂ x₁ + Y .dist y₂ y₁ 39 | ∎ 40 | where open ≤-Reasoning 41 | (X ⊗ Y) .triangle {x₁ , y₁} {x₂ , y₂} {x₃ , y₃} = 42 | begin 43 | X .dist x₁ x₃ + Y .dist y₁ y₃ 44 | ≤⟨ +-mono-≤ (X .triangle) (Y .triangle) ⟩ 45 | (X .dist x₁ x₂ + X .dist x₂ x₃) + (Y .dist y₁ y₂ + Y .dist y₂ y₃) 46 | ≈⟨ +-assoc _ _ _ ⟩ 47 | X .dist x₁ x₂ + (X .dist x₂ x₃ + (Y .dist y₁ y₂ + Y .dist y₂ y₃)) 48 | ≈⟨ +-cong ≃-refl (≃-sym (+-assoc _ _ _)) ⟩ 49 | X .dist x₁ x₂ + ((X .dist x₂ x₃ + Y .dist y₁ y₂) + Y .dist y₂ y₃) 50 | ≈⟨ +-cong ≃-refl (+-cong (+-comm _ _) ≃-refl) ⟩ 51 | X .dist x₁ x₂ + ((Y .dist y₁ y₂ + X .dist x₂ x₃) + Y .dist y₂ y₃) 52 | ≈⟨ +-cong ≃-refl (+-assoc _ _ _) ⟩ 53 | X .dist x₁ x₂ + (Y .dist y₁ y₂ + (X .dist x₂ x₃ + Y .dist y₂ y₃)) 54 | ≈⟨ ≃-sym (+-assoc _ _ _) ⟩ 55 | (X .dist x₁ x₂ + Y .dist y₁ y₂) + (X .dist x₂ x₃ + Y .dist y₂ y₃) 56 | ∎ 57 | where open ≤-Reasoning 58 | 59 | ≈-⊗ : ∀ {X Y} {x₁ x₂ : X .Carrier} {y₁ y₂ : Y .Carrier} → 60 | _≈_ X x₁ x₂ → _≈_ Y y₁ y₂ → _≈_ (X ⊗ Y) (x₁ , y₁) (x₂ , y₂) 61 | ≈-⊗ {X}{Y}{x₁}{x₂}{y₁}{y₂} x₁≈x₂ y₁≈y₂ = 62 | begin 63 | (X ⊗ Y) .dist (x₁ , y₁) (x₂ , y₂) 64 | ≡⟨⟩ 65 | X .dist x₁ x₂ + Y .dist y₁ y₂ 66 | ≤⟨ +-mono-≤ x₁≈x₂ y₁≈y₂ ⟩ 67 | 0ℝ + 0ℝ 68 | ≈⟨ +-identityʳ 0ℝ ⟩ 69 | 0ℝ 70 | ∎ 71 | where open ≤-Reasoning 72 | 73 | _⊗f_ : ∀ {X X' Y Y'} → X ⇒ X' → Y ⇒ Y' → (X ⊗ Y) ⇒ (X' ⊗ Y') 74 | (f ⊗f g) .fun (x , y) = f .fun x , g .fun y 75 | (f ⊗f g) .non-expansive {x₁ , y₁} {x₂ , y₂} = +-mono-≤ (f .non-expansive) (g .non-expansive) 76 | 77 | ⊗f-cong : ∀ {X X' Y Y'} {f₁ f₂ : X ⇒ X'}{g₁ g₂ : Y ⇒ Y'} → 78 | f₁ ≈f f₂ → g₁ ≈f g₂ → (f₁ ⊗f g₁) ≈f (f₂ ⊗f g₂) 79 | ⊗f-cong f₁≈f₂ g₁≈g₂ .f≈f (x , y) = 80 | ≤-trans (+-mono-≤ (f₁≈f₂ .f≈f x) (g₁≈g₂ .f≈f y)) (≤-reflexive (+-identityʳ 0ℝ)) 81 | 82 | -- FIXME: ⊗f preserves composition and identities 83 | 84 | ⊗f-∘ : ∀ {X₁ X₂ X₃ Y₁ Y₂ Y₃} → 85 | (f : X₂ ⇒ X₃) → (g : X₁ ⇒ X₂) → 86 | (f' : Y₂ ⇒ Y₃) → (g' : Y₁ ⇒ Y₂) → 87 | ((f ∘ g) ⊗f (f' ∘ g')) ≈f ((f ⊗f f') ∘ (g ⊗f g')) 88 | ⊗f-∘ {X₃ = X₃}{Y₃ = Y₃} f g f' g' .f≈f (x₁ , y₁) = 89 | (X₃ ⊗ Y₃) .≈-refl 90 | 91 | ------------------------------------------------------------------------------ 92 | -- From this product to the cartesian one. This could be derived from 93 | -- relationships between the norms used to define them. 94 | 95 | ⊗⇒× : ∀ {X Y} → (X ⊗ Y) ⇒ (X ×ₘ Y) 96 | ⊗⇒× .fun xy = xy 97 | ⊗⇒× .non-expansive = ⊔-least (+-increasingʳ _ _) (+-increasingˡ _ _) 98 | 99 | -- FIXME: this is natural, and preserves assoc, units, etc. 100 | 101 | ------------------------------------------------------------------------------ 102 | -- Associativity 103 | ⊗-assoc : ∀ {X Y Z} → (X ⊗ (Y ⊗ Z)) ⇒ ((X ⊗ Y) ⊗ Z) 104 | ⊗-assoc .fun (x , (y , z)) = ((x , y) , z) 105 | ⊗-assoc .non-expansive = ≤-reflexive (+-assoc _ _ _) 106 | 107 | ⊗-assoc⁻¹ : ∀ {X Y Z} → ((X ⊗ Y) ⊗ Z) ⇒ (X ⊗ (Y ⊗ Z)) 108 | ⊗-assoc⁻¹ .fun ((x , y) , z) = (x , (y , z)) 109 | ⊗-assoc⁻¹ .non-expansive = ≤-reflexive (≃-sym (+-assoc _ _ _)) 110 | 111 | -- FIXME: mutually inverse 112 | 113 | -- FIXME: pentagon 114 | 115 | ------------------------------------------------------------------------------ 116 | -- Swapping 117 | ⊗-symmetry : ∀ {X Y} → (X ⊗ Y) ⇒ (Y ⊗ X) 118 | ⊗-symmetry .fun (x , y) = y , x 119 | ⊗-symmetry .non-expansive = ≤-reflexive (+-comm _ _) 120 | 121 | ------------------------------------------------------------------------------ 122 | -- Units 123 | 124 | left-unit : ∀ {X} → (⊤ₘ ⊗ X) ⇒ X 125 | left-unit .fun (tt , x) = x 126 | left-unit .non-expansive = ≤-reflexive (≃-sym (+-identityˡ _)) 127 | 128 | left-unit⁻¹ : ∀ {X} → X ⇒ (⊤ₘ ⊗ X) 129 | left-unit⁻¹ .fun x = (tt , x) 130 | left-unit⁻¹ .non-expansive = ≤-reflexive (+-identityˡ _) 131 | 132 | module _ where 133 | 134 | left-unit-iso₁ : ∀ {X} → (left-unit ∘ left-unit⁻¹) ≈f (id {X}) 135 | left-unit-iso₁ {X} .f≈f a = X .≈-refl 136 | 137 | left-unit-iso₂ : ∀ {X} → (left-unit⁻¹ ∘ left-unit) ≈f (id {⊤ₘ ⊗ X}) 138 | left-unit-iso₂ {X} .f≈f (tt , a) = ≈-⊗ {⊤ₘ} {X} (⊤ₘ .≈-refl) (X .≈-refl) 139 | 140 | -- FIXME: define right-unit 141 | 142 | -- interaction with associativity 143 | -------------------------------------------------------------------------------- /src/MetricSpace/Rationals.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --allow-unsolved-metas #-} 2 | 3 | module MetricSpace.Rationals where 4 | 5 | open import Data.Product using (proj₁; proj₂; _,_; Σ-syntax) 6 | open import Data.Rational.Unnormalised as ℚ using () renaming (ℚᵘ to ℚ; 0ℚᵘ to 0ℚ; 1ℚᵘ to 1ℚ) 7 | import Data.Rational.Unnormalised.Properties as ℚ 8 | open import Data.Sum using (inj₁; inj₂) 9 | open import MetricSpace 10 | open import Data.Real.UpperClosed 11 | 12 | open MSpc 13 | open _⇒_ 14 | open _≈f_ 15 | 16 | private 17 | dist-sym : ∀ x y → ℚ.∣ x ℚ.- y ∣ ℚ.≃ ℚ.∣ y ℚ.- x ∣ 18 | dist-sym x y = 19 | begin-equality 20 | ℚ.∣ x ℚ.- y ∣ 21 | ≈⟨ ℚ.≃-sym (ℚ.∣-p∣≃∣p∣ (x ℚ.- y)) ⟩ 22 | ℚ.∣ ℚ.- (x ℚ.- y) ∣ 23 | ≈⟨ ℚ.∣-∣-cong (⁻¹-anti-homo-∙ x (ℚ.- y)) ⟩ 24 | ℚ.∣ ℚ.- (ℚ.- y) ℚ.- x ∣ 25 | ≈⟨ ℚ.∣-∣-cong (ℚ.+-congˡ (ℚ.- x) (⁻¹-involutive y)) ⟩ 26 | ℚ.∣ y ℚ.- x ∣ 27 | ∎ 28 | where open ℚ.≤-Reasoning 29 | open import Algebra.Properties.AbelianGroup ℚ.+-0-abelianGroup 30 | 31 | ℚspc : MSpc 32 | ℚspc .Carrier = ℚ 33 | ℚspc .dist x y = rational ℚ.∣ x ℚ.- y ∣ 34 | ℚspc .refl {x} = 35 | begin 36 | rational ℚ.∣ x ℚ.- x ∣ 37 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-inverseʳ x)) ⟩ 38 | rational ℚ.∣ 0ℚ ∣ 39 | ≈⟨ rational-cong (ℚ.0≤p⇒∣p∣≃p ℚ.≤-refl) ⟩ 40 | rational 0ℚ 41 | ≈⟨ rational-0 ⟩ 42 | 0ℝ 43 | ∎ 44 | where open ≤-Reasoning 45 | ℚspc .sym {x}{y} = 46 | begin 47 | rational ℚ.∣ x ℚ.- y ∣ 48 | ≈⟨ rational-cong (dist-sym x y) ⟩ 49 | rational ℚ.∣ y ℚ.- x ∣ 50 | ∎ 51 | where open ≤-Reasoning 52 | ℚspc .triangle {x}{y}{z} = 53 | begin 54 | rational ℚ.∣ x ℚ.- z ∣ 55 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.+-identityʳ (x ℚ.- z)))) ⟩ 56 | rational (ℚ.∣ (x ℚ.- z) ℚ.+ 0ℚ ∣) 57 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.+-congʳ (x ℚ.- z) (ℚ.+-inverseˡ y)))) ⟩ 58 | rational (ℚ.∣ (x ℚ.- z) ℚ.+ (ℚ.- y ℚ.+ y) ∣) 59 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ-interchange x (ℚ.- z) (ℚ.- y) y)) ⟩ 60 | rational (ℚ.∣ (x ℚ.- y) ℚ.+ ((ℚ.- z) ℚ.+ y) ∣) 61 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (x ℚ.- y) (ℚ.+-comm (ℚ.- z) y))) ⟩ 62 | rational (ℚ.∣ (x ℚ.- y) ℚ.+ (y ℚ.- z) ∣) 63 | ≤⟨ rational-mono (ℚ.∣p+q∣≤∣p∣+∣q∣ (x ℚ.- y) (y ℚ.- z)) ⟩ 64 | rational (ℚ.∣ x ℚ.- y ∣ ℚ.+ ℚ.∣ y ℚ.- z ∣) 65 | ≈⟨ ≃-sym (rational-+ ℚ.∣ x ℚ.- y ∣ ℚ.∣ y ℚ.- z ∣ (ℚ.0≤∣p∣ (x ℚ.- y)) (ℚ.0≤∣p∣ (y ℚ.- z))) ⟩ 66 | rational ℚ.∣ x ℚ.- y ∣ + rational ℚ.∣ y ℚ.- z ∣ 67 | ∎ 68 | where open ≤-Reasoning 69 | 70 | ℚspc-≈ : ∀ {q r} → q ℚ.≃ r → _≈_ ℚspc q r 71 | ℚspc-≈ {q}{r} q≃r = 72 | begin 73 | rational ℚ.∣ q ℚ.- r ∣ 74 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congˡ (ℚ.- r) q≃r)) ⟩ 75 | rational ℚ.∣ r ℚ.- r ∣ 76 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-inverseʳ r)) ⟩ 77 | rational ℚ.∣ 0ℚ ∣ 78 | ≈⟨ rational-cong (ℚ.0≤p⇒∣p∣≃p ℚ.≤-refl) ⟩ 79 | rational 0ℚ 80 | ≈⟨ rational-0 ⟩ 81 | 0ℝ 82 | ∎ 83 | where open ≤-Reasoning 84 | 85 | open import MetricSpace.MonoidalProduct 86 | open import MetricSpace.Terminal 87 | open import MetricSpace.InternalHom 88 | open import Data.Rational.Unnormalised.Positive as ℚ⁺ using (ℚ⁺) 89 | open MetricSpace.category 90 | 91 | const : ℚ → ⊤ₘ ⇒ ℚspc 92 | const q .fun _ = q 93 | const q .non-expansive = ℚspc .refl {q} 94 | 95 | ------------------------------------------------------------------------------ 96 | -- ℚspc is a commutative monoid object in the category of metric 97 | -- spaces and non-expansive maps 98 | 99 | zero : ⊤ₘ ⇒ ℚspc 100 | zero = const 0ℚ 101 | 102 | add : (ℚspc ⊗ ℚspc) ⇒ ℚspc 103 | add .fun (a , b) = a ℚ.+ b 104 | add .non-expansive {a₁ , b₁}{a₂ , b₂} = 105 | begin 106 | rational ℚ.∣ a₁ ℚ.+ b₁ ℚ.- (a₂ ℚ.+ b₂) ∣ 107 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (a₁ ℚ.+ b₁) (ℚ.≃-sym (⁻¹-∙-comm a₂ b₂)))) ⟩ 108 | rational ℚ.∣ a₁ ℚ.+ b₁ ℚ.+ (ℚ.- a₂ ℚ.- b₂) ∣ 109 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ-interchange a₁ b₁ (ℚ.- a₂) (ℚ.- b₂))) ⟩ 110 | rational ℚ.∣ (a₁ ℚ.- a₂) ℚ.+ (b₁ ℚ.- b₂) ∣ 111 | ≤⟨ rational-mono (ℚ.∣p+q∣≤∣p∣+∣q∣ (a₁ ℚ.- a₂) (b₁ ℚ.- b₂)) ⟩ 112 | rational (ℚ.∣ a₁ ℚ.- a₂ ∣ ℚ.+ ℚ.∣ b₁ ℚ.- b₂ ∣) 113 | ≈⟨ ≃-sym (rational-+ ℚ.∣ a₁ ℚ.- a₂ ∣ ℚ.∣ b₁ ℚ.- b₂ ∣ (ℚ.0≤∣p∣ (a₁ ℚ.- a₂)) (ℚ.0≤∣p∣ (b₁ ℚ.- b₂))) ⟩ 114 | rational ℚ.∣ a₁ ℚ.- a₂ ∣ + rational ℚ.∣ b₁ ℚ.- b₂ ∣ 115 | ≈⟨ ≃-refl ⟩ 116 | dist (ℚspc ⊗ ℚspc) (a₁ , b₁) (a₂ , b₂) 117 | ∎ 118 | where open ≤-Reasoning 119 | open import Algebra.Properties.AbelianGroup (ℚ.+-0-abelianGroup) 120 | 121 | add-identityˡ : (add ∘ (zero ⊗f id) ∘ left-unit⁻¹) ≈f id 122 | add-identityˡ .f≈f a = ℚspc-≈ (ℚ.+-identityˡ a) 123 | 124 | add-comm : (add ∘ ⊗-symmetry) ≈f add 125 | add-comm .f≈f (a , b) = ℚspc-≈ (ℚ.+-comm b a) 126 | 127 | add-assoc : (add ∘ (add ⊗f id) ∘ ⊗-assoc) ≈f (add ∘ (id ⊗f add)) 128 | add-assoc .f≈f (a , (b , c)) = ℚspc-≈ (ℚ.+-assoc a b c) 129 | 130 | ------------------------------------------------------------------------------ 131 | -- Negation, so we have a (graded) abelian group object 132 | open import MetricSpace.Scaling 133 | 134 | negate : ℚspc ⇒ ℚspc 135 | negate .fun a = ℚ.- a 136 | negate .non-expansive {a}{b} = 137 | begin 138 | rational ℚ.∣ ℚ.- a ℚ.- ℚ.- b ∣ 139 | ≈⟨ rational-cong (ℚ.∣-∣-cong (⁻¹-∙-comm a (ℚ.- b))) ⟩ 140 | rational ℚ.∣ ℚ.- (a ℚ.- b) ∣ 141 | ≈⟨ rational-cong (ℚ.∣-p∣≃∣p∣ (a ℚ.- b)) ⟩ 142 | dist ℚspc a b 143 | ∎ 144 | where open ≤-Reasoning 145 | open import Algebra.Properties.AbelianGroup (ℚ.+-0-abelianGroup) 146 | 147 | -- ![ 2 ] ℚspc → ℚspc 148 | add-inverse : (add ∘ (id ⊗f negate) ∘ (derelict ⊗f derelict) ∘ duplicate) ≈f (zero ∘ discard) 149 | add-inverse .f≈f q = ℚspc-≈ (ℚ.+-inverseʳ q) 150 | 151 | 152 | ------------------------------------------------------------------------------ 153 | -- binary max 154 | 155 | open import MetricSpace.CartesianProduct 156 | 157 | max : (ℚspc ×ₘ ℚspc) ⇒ ℚspc 158 | max .fun (x , y) = x ℚ.⊔ y 159 | max .non-expansive {x₁ , y₁}{x₂ , y₂} = {!!} 160 | 161 | {- ((x₁ ⊔ y₁) - (x₂ ⊔ y₂)) ⊔ ((x₂ ⊔ y₂) - (x₁ ⊔ y₁)) 162 | ≤ 163 | | x₁ - x₂ | ⊔ | y₁ - y₂ | 164 | ≃ 165 | (x₁ - x₂) ⊔ (x₂ - x₁) ⊔ (y₁ - y₂) ⊔ (y₂ - y₁) 166 | 167 | -} 168 | 169 | 170 | 171 | {- | x₁ ⊔ y₁ - x₂ ⊔ y₂ | 172 | ≤ 173 | | x₁ - x₂ | + | y₁ - y₂ | 174 | 175 | if x₁ ≤ y₁ then 176 | Goal: | y₁ - x₂ ⊔ y₂ | ≤ | x₁ - x₂ | + | y₁ - y₂ | 177 | if x₂ ≤ y₂ then 178 | Goal: | y₁ - y₂ | ≤ | x₁ - x₂ | + | y₁ - y₂ | 179 | ok 180 | else 181 | Goal: | y₁ - x₂ | ≤ 182 | -} 183 | 184 | ------------------------------------------------------------------------------ 185 | -- Multiplication by a "scalar" 186 | 187 | -- FIXME: make this work for any q : ℚ; need scaling to work for non-negative rationals 188 | ℚ-scale : (q : ℚ⁺) → ![ q ] ℚspc ⇒ ℚspc 189 | ℚ-scale q .fun a = ℚ⁺.fog q ℚ.* a 190 | ℚ-scale q .non-expansive {a}{b} = 191 | begin 192 | rational ℚ.∣ ℚ⁺.fog q ℚ.* a ℚ.- ℚ⁺.fog q ℚ.* b ∣ 193 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (ℚ⁺.fog q ℚ.* a) (ℚ.neg-distribʳ-* (ℚ⁺.fog q) b))) ⟩ 194 | rational ℚ.∣ ℚ⁺.fog q ℚ.* a ℚ.+ ℚ⁺.fog q ℚ.* (ℚ.- b) ∣ 195 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.*-distribˡ-+ (ℚ⁺.fog q) a (ℚ.- b)))) ⟩ 196 | rational ℚ.∣ ℚ⁺.fog q ℚ.* (a ℚ.- b) ∣ 197 | ≈⟨ rational-cong (ℚ.∣p*q∣≃∣p∣*∣q∣ (ℚ⁺.fog q) (a ℚ.- b)) ⟩ 198 | rational (ℚ.∣ ℚ⁺.fog q ∣ ℚ.* ℚ.∣ a ℚ.- b ∣) 199 | ≈⟨ ≃-sym (rational-* ℚ.∣ ℚ⁺.fog q ∣ ℚ.∣ a ℚ.- b ∣ (ℚ.0≤∣p∣ (ℚ⁺.fog q)) (ℚ.0≤∣p∣ (a ℚ.- b))) ⟩ 200 | rational ℚ.∣ ℚ⁺.fog q ∣ * rational ℚ.∣ a ℚ.- b ∣ 201 | ≈⟨ *-cong (rational-cong (ℚ.0≤p⇒∣p∣≃p (ℚ⁺.fog-nonneg q))) ≃-refl ⟩ 202 | rational+ q * rational ℚ.∣ a ℚ.- b ∣ 203 | ≡⟨⟩ 204 | ![ q ] ℚspc .dist a b 205 | ∎ 206 | where open ≤-Reasoning 207 | 208 | -- FIXME: make this work for any q : ℚ; need scaling to work for non-negative rationals 209 | ℚ-scale2 : (q : ℚ)(ε : ℚ⁺) → ℚ.∣ q ∣ ℚ.≤ ℚ⁺.fog ε → ![ ε ] ℚspc ⇒ ℚspc 210 | ℚ-scale2 q ε ∣q∣≤ε .fun a = q ℚ.* a 211 | ℚ-scale2 q ε ∣q∣≤ε .non-expansive {a}{b} = 212 | begin 213 | rational ℚ.∣ q ℚ.* a ℚ.- q ℚ.* b ∣ 214 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (q ℚ.* a) (ℚ.neg-distribʳ-* q b))) ⟩ 215 | rational ℚ.∣ q ℚ.* a ℚ.+ q ℚ.* (ℚ.- b) ∣ 216 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.*-distribˡ-+ q a (ℚ.- b)))) ⟩ 217 | rational ℚ.∣ q ℚ.* (a ℚ.- b) ∣ 218 | ≈⟨ rational-cong (ℚ.∣p*q∣≃∣p∣*∣q∣ q (a ℚ.- b)) ⟩ 219 | rational (ℚ.∣ q ∣ ℚ.* ℚ.∣ (a ℚ.- b) ∣) 220 | ≤⟨ rational-mono (ℚ.*-monoˡ-≤-nonNeg (ℚ.∣-∣-nonNeg (a ℚ.- b)) ∣q∣≤ε) ⟩ 221 | rational (ℚ⁺.fog ε ℚ.* ℚ.∣ (a ℚ.- b) ∣) 222 | ≈⟨ ≃-sym (rational-* (ℚ⁺.fog ε) ℚ.∣ a ℚ.- b ∣ (ℚ⁺.fog-nonneg ε) (ℚ.0≤∣p∣ (a ℚ.- b))) ⟩ 223 | rational+ ε * rational ℚ.∣ (a ℚ.- b) ∣ 224 | ≡⟨⟩ 225 | dist (![ ε ] ℚspc) a b 226 | ∎ 227 | where open ≤-Reasoning 228 | 229 | ------------------------------------------------------------------------------ 230 | -- Bounded ℚspc 231 | 232 | ℚspc[_] : ℚ⁺ → MSpc 233 | ℚspc[ b ] .Carrier = Σ[ q ∈ ℚ ] (ℚ.∣ q ∣ ℚ.≤ ℚ⁺.fog b) 234 | ℚspc[ b ] .dist (x , _) (y , _) = rational ℚ.∣ x ℚ.- y ∣ 235 | ℚspc[ b ] .refl {x , _} = 236 | begin 237 | rational ℚ.∣ x ℚ.- x ∣ 238 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-inverseʳ x)) ⟩ 239 | rational ℚ.∣ 0ℚ ∣ 240 | ≈⟨ rational-cong (ℚ.0≤p⇒∣p∣≃p ℚ.≤-refl) ⟩ 241 | rational 0ℚ 242 | ≈⟨ rational-0 ⟩ 243 | 0ℝ 244 | ∎ 245 | where open ≤-Reasoning 246 | ℚspc[ b ] .sym {x , _}{y , _} = 247 | begin 248 | rational ℚ.∣ x ℚ.- y ∣ 249 | ≈⟨ rational-cong (ℚ.≃-sym (ℚ.∣-p∣≃∣p∣ (x ℚ.- y))) ⟩ 250 | rational ℚ.∣ ℚ.- (x ℚ.- y) ∣ 251 | ≈⟨ rational-cong (ℚ.∣-∣-cong (⁻¹-anti-homo-∙ x (ℚ.- y))) ⟩ 252 | rational ℚ.∣ ℚ.- (ℚ.- y) ℚ.- x ∣ 253 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congˡ (ℚ.- x) (⁻¹-involutive y))) ⟩ 254 | rational ℚ.∣ y ℚ.- x ∣ 255 | ∎ 256 | where open ≤-Reasoning 257 | open import Algebra.Properties.AbelianGroup ℚ.+-0-abelianGroup 258 | ℚspc[ b ] .triangle {x , _}{y , _}{z , _} = 259 | begin 260 | rational ℚ.∣ x ℚ.- z ∣ 261 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.+-identityʳ (x ℚ.- z)))) ⟩ 262 | rational (ℚ.∣ (x ℚ.- z) ℚ.+ 0ℚ ∣) 263 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.+-congʳ (x ℚ.- z) (ℚ.+-inverseˡ y)))) ⟩ 264 | rational (ℚ.∣ (x ℚ.- z) ℚ.+ (ℚ.- y ℚ.+ y) ∣) 265 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ-interchange x (ℚ.- z) (ℚ.- y) y)) ⟩ 266 | rational (ℚ.∣ (x ℚ.- y) ℚ.+ ((ℚ.- z) ℚ.+ y) ∣) 267 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (x ℚ.- y) (ℚ.+-comm (ℚ.- z) y))) ⟩ 268 | rational (ℚ.∣ (x ℚ.- y) ℚ.+ (y ℚ.- z) ∣) 269 | ≤⟨ rational-mono (ℚ.∣p+q∣≤∣p∣+∣q∣ (x ℚ.- y) (y ℚ.- z)) ⟩ 270 | rational (ℚ.∣ x ℚ.- y ∣ ℚ.+ ℚ.∣ y ℚ.- z ∣) 271 | ≈⟨ ≃-sym (rational-+ ℚ.∣ x ℚ.- y ∣ ℚ.∣ y ℚ.- z ∣ (ℚ.0≤∣p∣ (x ℚ.- y)) (ℚ.0≤∣p∣ (y ℚ.- z))) ⟩ 272 | rational ℚ.∣ x ℚ.- y ∣ + rational ℚ.∣ y ℚ.- z ∣ 273 | ∎ 274 | where open ≤-Reasoning 275 | 276 | unbound : ∀ b → ℚspc[ b ] ⇒ ℚspc 277 | unbound b .fun (q , _) = q 278 | unbound b .non-expansive = ≤-refl 279 | 280 | open import Data.Sum using (inj₁; inj₂) 281 | 282 | lemma : ∀ q b → (ℚ.- (ℚ⁺.fog b)) ℚ.≤ q → q ℚ.≤ ℚ⁺.fog b → ℚ.∣ q ∣ ℚ.≤ ℚ⁺.fog b 283 | lemma q b -b≤q q≤b with ℚ.∣p∣≡p∨∣p∣≡-p q 284 | ... | inj₁ ∣q∣≡q = 285 | begin 286 | ℚ.∣ q ∣ ≡⟨ ∣q∣≡q ⟩ 287 | q ≤⟨ q≤b ⟩ 288 | ℚ⁺.fog b ∎ 289 | where open ℚ.≤-Reasoning 290 | ... | inj₂ ∣q∣≡-q = 291 | begin 292 | ℚ.∣ q ∣ ≡⟨ ∣q∣≡-q ⟩ 293 | ℚ.- q ≤⟨ ℚ.neg-mono-≤ -b≤q ⟩ 294 | ℚ.- (ℚ.- ℚ⁺.fog b) ≈⟨ ℚ.neg-involutive (ℚ⁺.fog b) ⟩ 295 | ℚ⁺.fog b ∎ 296 | where open ℚ.≤-Reasoning 297 | 298 | lemma2 : ∀ q → 0ℚ ℚ.≤ q → ℚ.- q ℚ.≤ q 299 | lemma2 q 0≤q = 300 | begin 301 | ℚ.- q ≤⟨ ℚ.neg-mono-≤ 0≤q ⟩ 302 | ℚ.- (ℚ.- 0ℚ) ≈⟨ ℚ.neg-involutive 0ℚ ⟩ 303 | 0ℚ ≤⟨ 0≤q ⟩ 304 | q ∎ 305 | where open ℚ.≤-Reasoning 306 | 307 | open import Relation.Nullary 308 | 309 | clamp-value : ℚ → ℚ⁺ → ℚ 310 | clamp-value q b = (q ℚ.⊔ ℚ.- ℚ⁺.fog b) ℚ.⊓ ℚ⁺.fog b 311 | 312 | module clamping where 313 | 314 | open import Relation.Binary.PropositionalEquality using (refl) 315 | 316 | data ClampView (b : ℚ⁺) (q : ℚ) : Set where 317 | lower : q ℚ.≤ ℚ.- ℚ⁺.fog b → ClampView b q 318 | within : ℚ.∣ q ∣ ℚ.≤ ℚ⁺.fog b → ClampView b q 319 | higher : ℚ⁺.fog b ℚ.≤ q → ClampView b q 320 | 321 | clamp-view : ∀ b q → ClampView b q 322 | clamp-view b q with ℚ⁺.fog b ℚ.≤? q 323 | ... | yes b≤q = higher b≤q 324 | ... | no ¬b≤q with q ℚ.≤? ℚ.- ℚ⁺.fog b 325 | ... | yes q≤-b = lower q≤-b 326 | ... | no ¬q≤-b with ℚ.∣p∣≡p∨∣p∣≡-p q 327 | ... | inj₁ ∣q∣≡q = 328 | within (begin 329 | ℚ.∣ q ∣ 330 | ≡⟨ ∣q∣≡q ⟩ 331 | q 332 | <⟨ ℚ.≰⇒> ¬b≤q ⟩ 333 | ℚ⁺.fog b 334 | ∎) 335 | where open ℚ.≤-Reasoning 336 | ... | inj₂ ∣q∣≡-q = 337 | within (begin 338 | ℚ.∣ q ∣ 339 | ≡⟨ ∣q∣≡-q ⟩ 340 | ℚ.- q 341 | <⟨ ℚ.neg-mono-< (ℚ.≰⇒> ¬q≤-b) ⟩ 342 | ℚ.- ℚ.- ℚ⁺.fog b 343 | ≈⟨ ℚ.neg-involutive (ℚ⁺.fog b) ⟩ 344 | ℚ⁺.fog b 345 | ∎) 346 | where open ℚ.≤-Reasoning 347 | 348 | clamp-v : ℚ⁺ → ℚ → ℚ 349 | clamp-v b q with clamp-view b q 350 | ... | lower _ = ℚ.- ℚ⁺.fog b 351 | ... | within _ = q 352 | ... | higher _ = ℚ⁺.fog b 353 | 354 | clamp-lemma : ∀ b q → ℚ.∣ clamp-v b q ∣ ℚ.≤ ℚ⁺.fog b 355 | clamp-lemma b q with clamp-view b q 356 | ... | lower q≤-b = begin ℚ.∣ ℚ.- ℚ⁺.fog b ∣ ≈⟨ ℚ.∣-p∣≃∣p∣ (ℚ⁺.fog b) ⟩ ℚ.∣ ℚ⁺.fog b ∣ ≈⟨ ℚ⁺.∣-∣-fog b ⟩ ℚ⁺.fog b ∎ 357 | where open ℚ.≤-Reasoning 358 | ... | within ∣q∣≤b = ∣q∣≤b 359 | ... | higher b≤q = begin ℚ.∣ ℚ⁺.fog b ∣ ≈⟨ ℚ⁺.∣-∣-fog b ⟩ ℚ⁺.fog b ∎ 360 | where open ℚ.≤-Reasoning 361 | 362 | {- 363 | clamp : ∀ b → ℚspc ⇒ ℚspc[ b ] 364 | clamp b .fun q = clamp-v b q , clamp-lemma b q 365 | clamp b .non-expansive {q₁}{q₂} with clamp-view b q₁ 366 | clamp b .non-expansive {q₁}{q₂} | lower q₁≤-b with clamp-view b q₂ 367 | clamp b .non-expansive {q₁}{q₂} | lower q₁≤-b | lower q₂≤-b = {!!} 368 | clamp b .non-expansive {q₁}{q₂} | lower q₁≤-b | within ∣q₂∣≤b = {!!} 369 | clamp b .non-expansive {q₁}{q₂} | lower q₁≤-b | higher b≤q₂ = {!!} 370 | clamp b .non-expansive {q₁}{q₂} | within ∣q₁∣≤b with clamp-view b q₂ 371 | clamp b .non-expansive {q₁}{q₂} | within ∣q₁∣≤b | lower q₂≤-b = {!!} 372 | clamp b .non-expansive {q₁}{q₂} | within ∣q₁∣≤b | within ∣q₂∣≤b = ≤-refl 373 | clamp b .non-expansive {q₁}{q₂} | within ∣q₁∣≤b | higher b≤q₂ = {!!} 374 | clamp b .non-expansive {q₁}{q₂} | higher b≤q₁ with clamp-view b q₂ 375 | clamp b .non-expansive {q₁}{q₂} | higher b≤q₁ | lower q₂≤-b = {!!} 376 | clamp b .non-expansive {q₁}{q₂} | higher b≤q₁ | within ∣q₂∣≤b = {!!} 377 | clamp b .non-expansive {q₁}{q₂} | higher b≤q₁ | higher b≤q₂ = {!!} 378 | -} 379 | 380 | ------------------------------------------------------------------------------ 381 | -- Multiplication 382 | 383 | multiply : ∀ a b → ![ b ] (ℚspc[ a ]) ⇒ (![ a ] (ℚspc[ b ]) ⊸ ℚspc[ a ℚ⁺.* b ]) 384 | multiply a b .fun (q , ∣q∣≤a) .fun (r , ∣r∣≤b) = q ℚ.* r , q*r-bounded 385 | where open ℚ.≤-Reasoning 386 | q*r-bounded : ℚ.∣ q ℚ.* r ∣ ℚ.≤ ℚ⁺.fog (a ℚ⁺.* b) 387 | q*r-bounded = 388 | begin 389 | ℚ.∣ q ℚ.* r ∣ 390 | ≡⟨ ℚ.∣p*q∣≡∣p∣*∣q∣ q r ⟩ 391 | ℚ.∣ q ∣ ℚ.* ℚ.∣ r ∣ 392 | ≤⟨ ℚ.*-monoʳ-≤-nonNeg {ℚ.∣ q ∣} (ℚ.∣-∣-nonNeg q) ∣r∣≤b ⟩ 393 | ℚ.∣ q ∣ ℚ.* ℚ⁺.fog b 394 | ≤⟨ ℚ.*-monoˡ-≤-nonNeg (ℚ.nonNegative (ℚ⁺.fog-nonneg b)) ∣q∣≤a ⟩ 395 | ℚ⁺.fog a ℚ.* ℚ⁺.fog b 396 | ≈⟨ ℚ.≃-sym (ℚ⁺.*-fog a b) ⟩ 397 | ℚ⁺.fog (a ℚ⁺.* b) 398 | ∎ 399 | multiply a b .fun (q , ∣q∣≤a) .non-expansive {r₁ , ∣r₁∣≤b} {r₂ , ∣r₂∣≤b} = 400 | begin 401 | rational ℚ.∣ q ℚ.* r₁ ℚ.- q ℚ.* r₂ ∣ 402 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (q ℚ.* r₁) (ℚ.neg-distribʳ-* q r₂))) ⟩ 403 | rational ℚ.∣ q ℚ.* r₁ ℚ.+ q ℚ.* (ℚ.- r₂) ∣ 404 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.*-distribˡ-+ q r₁ (ℚ.- r₂)))) ⟩ 405 | rational ℚ.∣ q ℚ.* (r₁ ℚ.- r₂) ∣ 406 | ≈⟨ rational-cong (ℚ.∣p*q∣≃∣p∣*∣q∣ q (r₁ ℚ.- r₂)) ⟩ 407 | rational (ℚ.∣ q ∣ ℚ.* ℚ.∣ r₁ ℚ.- r₂ ∣) 408 | ≈⟨ ≃-sym (rational-* ℚ.∣ q ∣ ℚ.∣ r₁ ℚ.- r₂ ∣ (ℚ.0≤∣p∣ q) (ℚ.0≤∣p∣ (r₁ ℚ.- r₂))) ⟩ 409 | rational (ℚ.∣ q ∣) * rational ℚ.∣ r₁ ℚ.- r₂ ∣ 410 | ≤⟨ *-mono-≤ (rational-mono ∣q∣≤a) ≤-refl ⟩ 411 | rational+ a * rational ℚ.∣ r₁ ℚ.- r₂ ∣ 412 | ∎ 413 | where open ≤-Reasoning 414 | multiply a b .non-expansive {q₁ , ∣q₁∣≤a}{q₂ , ∣q₂∣≤a} = 415 | sup-least λ { (r , ∣r∣≤b) → 416 | begin 417 | rational ℚ.∣ q₁ ℚ.* r ℚ.- q₂ ℚ.* r ∣ 418 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (q₁ ℚ.* r) (ℚ.neg-distribˡ-* q₂ r))) ⟩ 419 | rational ℚ.∣ q₁ ℚ.* r ℚ.+ (ℚ.- q₂) ℚ.* r ∣ 420 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.*-distribʳ-+ r q₁ (ℚ.- q₂)))) ⟩ 421 | rational ℚ.∣ (q₁ ℚ.- q₂) ℚ.* r ∣ 422 | ≈⟨ rational-cong (ℚ.∣p*q∣≃∣p∣*∣q∣ (q₁ ℚ.- q₂) r) ⟩ 423 | rational (ℚ.∣ q₁ ℚ.- q₂ ∣ ℚ.* ℚ.∣ r ∣) 424 | ≈⟨ rational-cong (ℚ.*-comm ℚ.∣ q₁ ℚ.- q₂ ∣ ℚ.∣ r ∣) ⟩ 425 | rational (ℚ.∣ r ∣ ℚ.* ℚ.∣ q₁ ℚ.- q₂ ∣) 426 | ≈⟨ ≃-sym (rational-* ℚ.∣ r ∣ ℚ.∣ q₁ ℚ.- q₂ ∣ (ℚ.0≤∣p∣ r) (ℚ.0≤∣p∣ (q₁ ℚ.- q₂))) ⟩ 427 | rational (ℚ.∣ r ∣) * rational ℚ.∣ q₁ ℚ.- q₂ ∣ 428 | ≤⟨ *-mono-≤ (rational-mono ∣r∣≤b) ≤-refl ⟩ 429 | rational+ b * rational ℚ.∣ q₁ ℚ.- q₂ ∣ 430 | ∎ } 431 | where open ≤-Reasoning 432 | 433 | mul : ∀ a b → (![ b ] (ℚspc[ a ]) ⊗ ![ a ] (ℚspc[ b ])) ⇒ ℚspc[ a ℚ⁺.* b ] 434 | mul a b = Λ⁻¹ (multiply a b) 435 | 436 | -- FIXME: work out how to state the usual equations for multiplication 437 | -- with this sensitivity typing. 438 | 439 | ------------------------------------------------------------------------------ 440 | -- Reciprocals 441 | 442 | ℚspc[_⟩ : ℚ⁺ → MSpc 443 | ℚspc[ b ⟩ .Carrier = Σ[ q ∈ ℚ ] (ℚ⁺.fog b ℚ.≤ q) 444 | ℚspc[ b ⟩ .dist (x , _) (y , _) = rational ℚ.∣ x ℚ.- y ∣ 445 | ℚspc[ b ⟩ .refl {x , _} = 446 | begin 447 | rational ℚ.∣ x ℚ.- x ∣ 448 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-inverseʳ x)) ⟩ 449 | rational ℚ.∣ 0ℚ ∣ 450 | ≈⟨ rational-cong (ℚ.0≤p⇒∣p∣≃p ℚ.≤-refl) ⟩ 451 | rational 0ℚ 452 | ≈⟨ rational-0 ⟩ 453 | 0ℝ 454 | ∎ 455 | where open ≤-Reasoning 456 | ℚspc[ b ⟩ .sym {x , _}{y , _} = 457 | begin 458 | rational ℚ.∣ x ℚ.- y ∣ 459 | ≈⟨ rational-cong (ℚ.≃-sym (ℚ.∣-p∣≃∣p∣ (x ℚ.- y))) ⟩ 460 | rational ℚ.∣ ℚ.- (x ℚ.- y) ∣ 461 | ≈⟨ rational-cong (ℚ.∣-∣-cong (⁻¹-anti-homo-∙ x (ℚ.- y))) ⟩ 462 | rational ℚ.∣ ℚ.- (ℚ.- y) ℚ.- x ∣ 463 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congˡ (ℚ.- x) (⁻¹-involutive y))) ⟩ 464 | rational ℚ.∣ y ℚ.- x ∣ 465 | ∎ 466 | where open ≤-Reasoning 467 | open import Algebra.Properties.AbelianGroup ℚ.+-0-abelianGroup 468 | ℚspc[ b ⟩ .triangle {x , _}{y , _}{z , _} = 469 | begin 470 | rational ℚ.∣ x ℚ.- z ∣ 471 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.+-identityʳ (x ℚ.- z)))) ⟩ 472 | rational (ℚ.∣ (x ℚ.- z) ℚ.+ 0ℚ ∣) 473 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.+-congʳ (x ℚ.- z) (ℚ.+-inverseˡ y)))) ⟩ 474 | rational (ℚ.∣ (x ℚ.- z) ℚ.+ (ℚ.- y ℚ.+ y) ∣) 475 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ-interchange x (ℚ.- z) (ℚ.- y) y)) ⟩ 476 | rational (ℚ.∣ (x ℚ.- y) ℚ.+ ((ℚ.- z) ℚ.+ y) ∣) 477 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-congʳ (x ℚ.- y) (ℚ.+-comm (ℚ.- z) y))) ⟩ 478 | rational (ℚ.∣ (x ℚ.- y) ℚ.+ (y ℚ.- z) ∣) 479 | ≤⟨ rational-mono (ℚ.∣p+q∣≤∣p∣+∣q∣ (x ℚ.- y) (y ℚ.- z)) ⟩ 480 | rational (ℚ.∣ x ℚ.- y ∣ ℚ.+ ℚ.∣ y ℚ.- z ∣) 481 | ≈⟨ ≃-sym (rational-+ ℚ.∣ x ℚ.- y ∣ ℚ.∣ y ℚ.- z ∣ (ℚ.0≤∣p∣ (x ℚ.- y)) (ℚ.0≤∣p∣ (y ℚ.- z))) ⟩ 482 | rational ℚ.∣ x ℚ.- y ∣ + rational ℚ.∣ y ℚ.- z ∣ 483 | ∎ 484 | where open ≤-Reasoning 485 | 486 | module _ where 487 | open import Data.Integer as ℤ 488 | 489 | nonzero : ∀ a q → ℚ⁺.fog a ℚ.≤ q → ℤ.∣ ℚ.↥ q ∣ ℚ.≢0 490 | nonzero a q a≤q = ℚ.p≄0⇒∣↥p∣≢0 q λ q≃0 → 491 | ℚ⁺.fog-not≤0 a (begin ℚ⁺.fog a ≤⟨ a≤q ⟩ q ≈⟨ q≃0 ⟩ 0ℚ ∎) 492 | where open ℚ.≤-Reasoning 493 | 494 | recip : ∀ a → ![ ℚ⁺.1/ (a ℚ⁺.* a) ] ℚspc[ a ⟩ ⇒ ℚspc 495 | recip a .fun (q , a≤q) = ℚ.1/_ q {nonzero a q a≤q} 496 | recip a .non-expansive {q , a≤q} {r , a≤r} = 497 | begin 498 | rational ℚ.∣ ℚ.1/ q ℚ.- ℚ.1/ r ∣ 499 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-cong (ℚ.≃-sym (ℚ.*-identityˡ (ℚ.1/ q))) (ℚ.-‿cong (ℚ.≃-sym (ℚ.*-identityˡ (ℚ.1/ r)))))) ⟩ 500 | rational ℚ.∣ (1ℚ ℚ.* ℚ.1/ q) ℚ.- (1ℚ ℚ.* ℚ.1/ r) ∣ 501 | ≈⟨ ≃-sym (rational-cong (ℚ.∣-∣-cong (ℚ.+-cong (ℚ.*-cong (ℚ.*-inverseʳ r {r≢0}) (ℚ.≃-refl {ℚ.1/ q})) 502 | (ℚ.-‿cong (ℚ.*-cong (ℚ.*-inverseʳ q {q≢0}) (ℚ.≃-refl {ℚ.1/ r})))))) ⟩ 503 | rational ℚ.∣ ((r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q) ℚ.- ((q ℚ.* ℚ.1/ q) ℚ.* ℚ.1/ r) ∣ 504 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-cong (ℚ.≃-refl {(r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q}) 505 | (ℚ.-‿cong (ℚ.*-assoc q (ℚ.1/ q) (ℚ.1/ r))))) ⟩ 506 | rational ℚ.∣ ((r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q) ℚ.- (q ℚ.* (ℚ.1/ q ℚ.* ℚ.1/ r)) ∣ 507 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-cong (ℚ.≃-refl {(r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q}) 508 | (ℚ.-‿cong (ℚ.*-cong (ℚ.≃-refl {q}) (ℚ.*-comm (ℚ.1/ q) (ℚ.1/ r)))))) ⟩ 509 | rational ℚ.∣ ((r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q) ℚ.- (q ℚ.* (ℚ.1/ r ℚ.* ℚ.1/ q)) ∣ 510 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-cong (ℚ.≃-refl {(r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q}) 511 | (ℚ.neg-distribˡ-* q (ℚ.1/ r ℚ.* ℚ.1/ q)))) ⟩ 512 | rational ℚ.∣ ((r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q) ℚ.+ (ℚ.- q ℚ.* (ℚ.1/ r ℚ.* ℚ.1/ q)) ∣ 513 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.+-cong (ℚ.≃-refl {(r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q}) (ℚ.≃-sym (ℚ.*-assoc (ℚ.- q) (ℚ.1/ r) (ℚ.1/ q))))) ⟩ 514 | rational ℚ.∣ ((r ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q) ℚ.+ ((ℚ.- q ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q) ∣ 515 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.*-distribʳ-+ (ℚ.1/ q) (r ℚ.* ℚ.1/ r) (ℚ.- q ℚ.* ℚ.1/ r)))) ⟩ 516 | rational ℚ.∣ (r ℚ.* ℚ.1/ r ℚ.+ (ℚ.- q ℚ.* ℚ.1/ r)) ℚ.* ℚ.1/ q ∣ 517 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ.*-cong (ℚ.*-distribʳ-+ (ℚ.1/ r) r (ℚ.- q)) (ℚ.≃-refl {ℚ.1/ q})))) ⟩ 518 | rational ℚ.∣ ((r ℚ.- q) ℚ.* ℚ.1/ r) ℚ.* ℚ.1/ q ∣ 519 | ≈⟨ rational-cong (ℚ.∣-∣-cong (ℚ.*-assoc (r ℚ.- q) (ℚ.1/ r) (ℚ.1/ q))) ⟩ 520 | rational ℚ.∣ (r ℚ.- q) ℚ.* (ℚ.1/ r ℚ.* ℚ.1/ q) ∣ 521 | ≈⟨ rational-cong (ℚ.∣p*q∣≃∣p∣*∣q∣ (r ℚ.- q) (ℚ.1/ r ℚ.* ℚ.1/ q)) ⟩ 522 | rational (ℚ.∣ r ℚ.- q ∣ ℚ.* ℚ.∣ ℚ.1/ r ℚ.* ℚ.1/ q ∣) 523 | ≈⟨ rational-cong (ℚ.*-comm ℚ.∣ r ℚ.- q ∣ ℚ.∣ ℚ.1/ r ℚ.* ℚ.1/ q ∣) ⟩ 524 | rational (ℚ.∣ ℚ.1/ r ℚ.* ℚ.1/ q ∣ ℚ.* ℚ.∣ r ℚ.- q ∣) 525 | ≈⟨ ≃-sym (rational-* _ _ (ℚ.0≤∣p∣ (ℚ.1/ r ℚ.* ℚ.1/ q)) (ℚ.0≤∣p∣ (r ℚ.- q))) ⟩ 526 | rational ℚ.∣ ℚ.1/ r ℚ.* ℚ.1/ q ∣ * rational ℚ.∣ r ℚ.- q ∣ 527 | ≈⟨ ≃-refl ⟩ 528 | rational ℚ.∣ ℚ.1/ (ℚ⁺.fog r⁺) ℚ.* ℚ.1/ (ℚ⁺.fog q⁺) ∣ * rational ℚ.∣ r ℚ.- q ∣ 529 | ≈⟨ *-cong (rational-cong (ℚ.∣-∣-cong (ℚ.*-cong (ℚ⁺.1/-fog r⁺ r≢0) (ℚ⁺.1/-fog q⁺ q≢0)))) ≃-refl ⟩ 530 | rational ℚ.∣ ℚ⁺.fog (ℚ⁺.1/ r⁺) ℚ.* ℚ⁺.fog (ℚ⁺.1/ q⁺) ∣ * rational ℚ.∣ r ℚ.- q ∣ 531 | ≈⟨ *-cong (rational-cong (ℚ.∣-∣-cong (ℚ.≃-sym (ℚ⁺.*-fog (ℚ⁺.1/ r⁺) (ℚ⁺.1/ q⁺))))) ≃-refl ⟩ 532 | rational ℚ.∣ ℚ⁺.fog (ℚ⁺.1/ r⁺ ℚ⁺.* ℚ⁺.1/ q⁺) ∣ * rational ℚ.∣ r ℚ.- q ∣ 533 | ≈⟨ *-cong (rational-cong (ℚ⁺.∣-∣-fog (ℚ⁺.1/ r⁺ ℚ⁺.* ℚ⁺.1/ q⁺))) ≃-refl ⟩ 534 | rational+ (ℚ⁺.1/ r⁺ ℚ⁺.* ℚ⁺.1/ q⁺) * rational ℚ.∣ r ℚ.- q ∣ 535 | ≤⟨ *-mono-≤ (rational-mono (ℚ⁺.fog-mono (ℚ⁺.*-mono-≤ (ℚ⁺.1/-invert-≤ a r⁺ (ℚ⁺.r≤r a≤r)) 536 | (ℚ⁺.1/-invert-≤ a q⁺ (ℚ⁺.r≤r a≤q))))) 537 | ≤-refl ⟩ 538 | rational+ (ℚ⁺.1/ a ℚ⁺.* ℚ⁺.1/ a) * rational ℚ.∣ r ℚ.- q ∣ 539 | ≈⟨ *-cong (rational-cong (ℚ⁺.fog-cong (⁻¹-∙-comm a a))) (rational-cong (dist-sym r q)) ⟩ 540 | rational+ (ℚ⁺.1/ (a ℚ⁺.* a)) * rational ℚ.∣ q ℚ.- r ∣ 541 | ∎ 542 | where open ≤-Reasoning 543 | open import Algebra.Properties.AbelianGroup ℚ⁺.*-1-AbelianGroup 544 | r≢0 = nonzero a r a≤r 545 | q≢0 = nonzero a q a≤q 546 | r⁺ = ℚ⁺.⟨ r , ℚ.positive (ℚ.<-≤-trans (ℚ.positive⁻¹ (a .ℚ⁺.positive)) a≤r) ⟩ 547 | q⁺ = ℚ⁺.⟨ q , ℚ.positive (ℚ.<-≤-trans (ℚ.positive⁻¹ (a .ℚ⁺.positive)) a≤q) ⟩ 548 | -------------------------------------------------------------------------------- /src/MetricSpace/Scaling.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module MetricSpace.Scaling where 4 | 5 | open import Data.Product using (_,_) 6 | open import Data.Real.UpperClosed 7 | open import Data.Rational.Unnormalised.Positive as ℚ⁺ using (ℚ⁺; 1ℚ⁺) 8 | open import MetricSpace 9 | open import MetricSpace.Category 10 | open import MetricSpace.MonoidalProduct 11 | 12 | open MSpc 13 | open _⇒_ 14 | open _≈f_ 15 | 16 | -- A graded comonad, for expressing sensitivity wrt inputs 17 | 18 | -- FIXME: ought to be ℚ≥0, or even ℝᵘ (but that isn't a set! is that a 19 | -- problem?) 20 | ![_] : ℚ⁺ → MSpc → MSpc 21 | ![ ε ] X .Carrier = X .Carrier 22 | ![ ε ] X .dist x y = rational+ ε * X .dist x y 23 | ![ ε ] X .refl {x} = 24 | begin 25 | rational+ ε * X .dist x x 26 | ≤⟨ *-mono-≤ ≤-refl (X .refl) ⟩ 27 | rational+ ε * 0ℝ 28 | ≈⟨ *-zeroʳ (rational+ ε) (rational+<∞ ε) ⟩ 29 | 0ℝ 30 | ∎ 31 | where open ≤-Reasoning 32 | ![ ε ] X .sym {x}{y} = 33 | begin 34 | rational+ ε * X .dist x y 35 | ≤⟨ *-mono-≤ ≤-refl (X .sym) ⟩ 36 | rational+ ε * X .dist y x 37 | ∎ 38 | where open ≤-Reasoning 39 | ![ ε ] X .triangle {x}{y}{z} = 40 | begin 41 | rational+ ε * X .dist x z 42 | ≤⟨ *-mono-≤ ≤-refl (X .triangle) ⟩ 43 | rational+ ε * (X .dist x y + X .dist y z) 44 | ≈⟨ *-distribˡ-+ (rational+ ε) (X .dist x y) (X .dist y z) ⟩ 45 | (rational+ ε * X .dist x y) + (rational+ ε * X .dist y z) 46 | ∎ 47 | where open ≤-Reasoning 48 | 49 | map : ∀ {ε X Y} → X ⇒ Y → ![ ε ] X ⇒ ![ ε ] Y 50 | map f .fun x = f .fun x 51 | map f .non-expansive = *-mono-≤ ≤-refl (f .non-expansive) 52 | 53 | map-cong : ∀ {ε X Y} {f₁ f₂ : X ⇒ Y} → f₁ ≈f f₂ → map {ε} f₁ ≈f map {ε} f₂ 54 | map-cong {ε}{X}{Y} {f₁}{f₂} f₁≈f₂ .f≈f x = 55 | begin 56 | rational+ ε * Y .dist (f₁ .fun x) (f₂ .fun x) 57 | ≤⟨ *-mono-≤ ≤-refl (f₁≈f₂ .f≈f x) ⟩ 58 | rational+ ε * 0ℝ 59 | ≈⟨ *-zeroʳ (rational+ ε) (rational+<∞ ε) ⟩ 60 | 0ℝ 61 | ∎ 62 | where open ≤-Reasoning 63 | 64 | 65 | ------------------------------------------------------------------------------ 66 | -- Functoriality properties 67 | 68 | map-id : ∀ {ε X} → map {ε}{X} id ≈f id 69 | map-id {ε}{X} .f≈f x = ![ ε ] X .refl {x} 70 | 71 | map-∘ : ∀ {ε X Y Z} {f : Y ⇒ Z} {g : X ⇒ Y} → 72 | map {ε} (f ∘ g) ≈f (map f ∘ map g) 73 | map-∘ {ε}{Z = Z}{f}{g} .f≈f a = ![ ε ] Z .refl {map {ε} f .fun (map {ε} g .fun a)} 74 | 75 | weaken : ∀ {ε₁ ε₂ X} → ε₂ ℚ⁺.≤ ε₁ → ![ ε₁ ] X ⇒ ![ ε₂ ] X 76 | weaken ε₂≤ε₁ .fun x = x 77 | weaken ε₂≤ε₁ .non-expansive = *-mono-≤ (rational-mono (ℚ⁺.fog-mono ε₂≤ε₁)) ≤-refl 78 | -- FIXME: weaken is natural, and functorial 79 | 80 | ------------------------------------------------------------------------------ 81 | -- ![_] is a graded exponential comonad with respect to the monoidal 82 | -- !product, so there are quite a few properties to prove for the 83 | -- !natural transformations defined below. 84 | 85 | -- FIXME: discard : ![ 0 ] X ⇒ ⊤ₘ -- can't do this yet since using ℚ⁺ 86 | 87 | 88 | derelict : ∀ {X} → ![ 1ℚ⁺ ] X ⇒ X 89 | derelict .fun x = x 90 | derelict {X} .non-expansive = ≤-reflexive (≃-sym (*-identityˡ (X .dist _ _))) 91 | -- FIXME: derelict is natural 92 | 93 | dig : ∀ {ε₁ ε₂ X} → ![ ε₁ ℚ⁺.* ε₂ ] X ⇒ ![ ε₁ ] (![ ε₂ ] X) 94 | dig .fun x = x 95 | dig {ε₁}{ε₂}{X} .non-expansive {a}{b} = 96 | begin 97 | rational+ ε₁ * (rational+ ε₂ * X .dist a b) 98 | ≈⟨ ≃-sym (*-assoc (rational+ ε₁) (rational+ ε₂) (X .dist a b)) ⟩ 99 | (rational+ ε₁ * rational+ ε₂) * X .dist a b 100 | ≈⟨ *-cong (rational⁺-* ε₁ ε₂) ≃-refl ⟩ 101 | rational+ (ε₁ ℚ⁺.* ε₂) * X .dist a b 102 | ∎ 103 | where open ≤-Reasoning 104 | -- FIXME: dig is natural 105 | 106 | duplicate : ∀ {ε₁ ε₂ X} → ![ ε₁ ℚ⁺.+ ε₂ ] X ⇒ (![ ε₁ ] X ⊗ ![ ε₂ ] X) 107 | duplicate .fun x = (x , x) 108 | duplicate {ε₁}{ε₂}{X} .non-expansive {a}{b} = 109 | begin 110 | (rational+ ε₁ * X .dist a b) + (rational+ ε₂ * X .dist a b) 111 | ≈⟨ ≃-sym (*-distribʳ-+ (X .dist a b) (rational+ ε₁) (rational+ ε₂)) ⟩ 112 | (rational+ ε₁ + rational+ ε₂) * X .dist a b 113 | ≈⟨ *-cong (rational⁺-+ ε₁ ε₂) ≃-refl ⟩ 114 | rational+ (ε₁ ℚ⁺.+ ε₂) * X .dist a b 115 | ≈⟨ ≃-refl ⟩ 116 | dist (![ ε₁ ℚ⁺.+ ε₂ ] X) a b 117 | ∎ 118 | where open ≤-Reasoning 119 | -- FIXME: duplicate is natural 120 | 121 | monoidal : ∀ {ε X Y} → (![ ε ] X ⊗ ![ ε ] Y) ⇒ ![ ε ] (X ⊗ Y) 122 | monoidal .fun (x , y) = x , y 123 | monoidal {ε}{X}{Y} .non-expansive {x₁ , y₁}{x₂ , y₂} = 124 | begin 125 | rational+ ε * (X .dist x₁ x₂ + Y .dist y₁ y₂) 126 | ≈⟨ *-distribˡ-+ (rational+ ε) (X .dist x₁ x₂) (Y .dist y₁ y₂) ⟩ 127 | (rational+ ε * X .dist x₁ x₂) + (rational+ ε * Y .dist y₁ y₂) 128 | ∎ 129 | where open ≤-Reasoning 130 | -- FIXME: this witnesses ![ ε ] as a monoidal functor: 131 | -- naturality 132 | -- commutes with assoc and swapping 133 | -------------------------------------------------------------------------------- /src/MetricSpace/Terminal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module MetricSpace.Terminal where 4 | 5 | open import Data.Unit using (⊤; tt) 6 | open import MetricSpace 7 | open import Data.Real.UpperClosed 8 | 9 | open MSpc 10 | open _⇒_ 11 | open _≈f_ 12 | 13 | ⊤ₘ : MSpc 14 | ⊤ₘ .Carrier = ⊤ 15 | ⊤ₘ .dist tt tt = 0ℝ 16 | ⊤ₘ .refl = ≤-refl 17 | ⊤ₘ .sym = ≤-refl 18 | ⊤ₘ .triangle = ≤-reflexive (≃-sym (+-identityˡ 0ℝ)) 19 | 20 | discard : ∀ {X} → X ⇒ ⊤ₘ 21 | discard .fun _ = tt 22 | discard .non-expansive = 0-least _ 23 | 24 | discard-unique : ∀ {X} {f : X ⇒ ⊤ₘ} → f ≈f discard 25 | discard-unique .f≈f a = ≤-refl 26 | --------------------------------------------------------------------------------