------------------------------------------------------------------------ -- Agda Helium -- -- Algebraic properties of types used by the Armv8-M pseudocode. ------------------------------------------------------------------------ {-# OPTIONS --safe --without-K #-} open import Helium.Data.Pseudocode.Algebra module Helium.Data.Pseudocode.Algebra.Properties {b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃} (pseudocode : Pseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃) where open import Agda.Builtin.FromNat open import Agda.Builtin.FromNeg open import Data.Nat using (ℕ; suc; NonZero) import Data.Nat.Literals as ℕₗ open import Data.Product as × using (∃; _×_; _,_) open import Data.Sum using (_⊎_; inj₁; inj₂; map) import Data.Unit open import Function using (_∘_) import Helium.Algebra.Ordered.StrictTotal.Properties.CommutativeRing as CommRingₚ import Helium.Algebra.Ordered.StrictTotal.Properties.Field as Fieldₚ open import Level using (_⊔_) open import Relation.Binary.Core using (_Preserves_⟶_) open import Relation.Binary.Definitions using (tri<; tri≈; tri>) open import Relation.Nullary.Negation using (contradiction) private module pseudocode = Pseudocode pseudocode instance numberℕ : Number ℕ numberℕ = ℕₗ.number open pseudocode public hiding ( integerDiscrete; 1≉0; ⌊⌋-floor ; module ℤ; module ℤ′; module ℤ-Reasoning ; module ℝ; module ℝ′; module ℝ-Reasoning ; module ⌊⌋ ; module /1 ) module ℤ where open pseudocode.ℤ public hiding (ℤ; 0ℤ; 1ℤ) renaming ( trans to <-trans ; irrefl to <-irrefl ; asym to <-asym ; 00∧y>0⇒xy>0 ) module Reasoning = pseudocode.ℤ-Reasoning open CommRingₚ integerRing public discrete : ∀ x y → y ≤ x ⊎ x + 1ℤ ≤ y discrete = pseudocode.integerDiscrete 1≉0 : 1ℤ ≉ 0ℤ 1≉0 = pseudocode.1≉0 x≤y0∧y>0⇒x*y>0 ) module Reasoning = pseudocode.ℝ-Reasoning open Fieldₚ realField public hiding () 1≉0 : 1ℝ ≉ 0ℝ 1≉0 1≈0 = ℤ.1≉0 (begin-equality 1ℤ ≈˘⟨ ⌊x/1⌋≈x 1ℤ ⟩ ⌊ 1ℤ /1 ⌋ ≈⟨ ⌊⌋.cong /1.1#-homo ⟩ ⌊ 1ℝ ⌋ ≈⟨ ⌊⌋.cong 1≈0 ⟩ ⌊ 0ℝ ⌋ ≈˘⟨ ⌊⌋.cong /1.0#-homo ⟩ ⌊ 0ℤ /1 ⌋ ≈⟨ ⌊x/1⌋≈x 0ℤ ⟩ 0ℤ ∎) where open ℤ.Reasoning open pseudocode using (module /1; module ⌊⌋) record IsMonotoneContinuous (f : ℝ → ℝ) : Set (r₁ ⊔ r₂ ⊔ r₃) where field cong : f Preserves _≈_ ⟶ _≈_ mono-≤ : f Preserves _≤_ ⟶ _≤_ continuous-≤-< : ∀ {x y z} → f x ≤ z → z < f y → ∃ λ α → f α ≈ z × x ≤ α × α < y continuous-<-≤ : ∀ {x y z} → f x < z → z ≤ f y → ∃ λ α → f α ≈ z × x < α × α ≤ y continuous-<-< : ∀ {x y z} → f x < z → z < f y → ∃ λ α → f α ≈ z × x < α × α < y record MCHelper (f : ℝ → ℝ) : Set (r₁ ⊔ r₂ ⊔ r₃) where field cong : f Preserves _≈_ ⟶ _≈_ mono-≤ : f Preserves _≤_ ⟶ _≤_ continuous-<-< : ∀ {x y z} → f x < z → z < f y → ∃ λ α → f α ≈ z × x < α × α < y mcHelper : ∀ {f} → MCHelper f → IsMonotoneContinuous f mcHelper {f} helper = record { cong = cong ; mono-≤ = mono-≤ ; continuous-≤-< = continuous-≤-< ; continuous-<-≤ = continuous-<-≤ ; continuous-<-< = continuous-<-< } where open MCHelper helper cancel-< : ∀ {x y} → f x < f y → x < y cancel-< fx (<⇒≱ fx (λ x≥y → ℤ.<⇒≱ ⌊x⌋<⌊y⌋ (mono-≤ x≥y)) floor : ∀ {x y} → x ℝ.< y /1 → ⌊ x ⌋ ℤ.< y floor {x} {y} = pseudocode.⌊⌋-floor x y n≤x⇒n≤⌊x⌋ : ∀ {n x} → n /1 ℝ.≤ x → n ℤ.≤ ⌊ x ⌋ n≤x⇒n≤⌊x⌋ {n} {x} n≤x = begin n ≈˘⟨ ⌊x/1⌋≈x n ⟩ ⌊ n /1 ⌋ ≤⟨ mono-≤ n≤x ⟩ ⌊ x ⌋ ∎ where open ℤ.Reasoning ⌊x⌋≤x : ∀ x → ⌊ x ⌋ /1 ℝ.≤ x ⌊x⌋≤x x = ℝ.≮⇒≥ (λ x<⌊x⌋ → ℤ.<-asym (floor x<⌊x⌋) (floor x<⌊x⌋)) x<⌊x⌋+1 : ∀ x → x ℝ.< (⌊ x ⌋ ℤ.+ 1ℤ) /1 x<⌊x⌋+1 x = ℝ.≰⇒> (λ x≥⌊x⌋+1 → ℤ.<⇒≱ (ℤ.≤∧≉⇒< ℤ.0≤1 (ℤ.1≉0 ∘ ℤ.Eq.sym)) (ℤ.+-cancelˡ-≤ ⌊ x ⌋ (begin ⌊ x ⌋ ℤ.+ 1ℤ ≤⟨ n≤x⇒n≤⌊x⌋ x≥⌊x⌋+1 ⟩ ⌊ x ⌋ ≈˘⟨ ℤ.+-identityʳ ⌊ x ⌋ ⟩ ⌊ x ⌋ ℤ.+ 0ℤ ∎))) where open ℤ.Reasoning n≤x