diff --git a/CHANGELOG.md b/CHANGELOG.md index c745d6be4f..93dc9468ab 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -808,6 +808,17 @@ Other minor changes finTofun-funToFin : funToFin ∘ finToFun ≗ id funTofin-funToFun : finToFun (funToFin f) ≗ f ^↔→ : Extensionality _ _ → Fin (n ^ m) ↔ (Fin m → Fin n) + + toℕ-mono-< : i < j → toℕ i ℕ.< toℕ j + toℕ-mono-≤ : i ≤ j → toℕ i ℕ.≤ toℕ j + toℕ-cancel-≤ : toℕ i ℕ.≤ toℕ j → i ≤ j + toℕ-cancel-< : toℕ i ℕ.< toℕ j → i < j + + toℕ-combine : toℕ (combine x y) ≡ k ℕ.* toℕ x ℕ.+ toℕ y + combine-injectiveˡ : combine x z ≡ combine y z → x ≡ y + combine-injectiveʳ : combine x y ≡ combine x z → y ≡ z + combine-injective : combine x y ≡ combine w z → x ≡ w × y ≡ z + combine-surjective : ∀ x → ∃₂ λ y z → combine y z ≡ x ``` * Added new functions in `Data.Integer.Base`: diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda index 2ba45794a8..b10bffd54c 100644 --- a/src/Data/Fin/Properties.agda +++ b/src/Data/Fin/Properties.agda @@ -19,8 +19,10 @@ open import Data.Fin.Base open import Data.Fin.Patterns open import Data.Nat.Base as ℕ using (ℕ; zero; suc; s≤s; z≤n; _∸_; _^_) import Data.Nat.Properties as ℕₚ +open import Data.Nat.Solver open import Data.Unit using (⊤; tt) open import Data.Product using (Σ-syntax; ∃; ∃₂; ∄; _×_; _,_; map; proj₁; proj₂; uncurry; <_,_>) +open import Data.Product.Properties using (,-injective) open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]; [_,_]′) open import Data.Sum.Properties using ([,]-map-commute; [,]-∘-distr) open import Function.Base using (_∘_; id; _$_; flip) @@ -28,7 +30,7 @@ open import Function.Bundles using (_↣_; _⇔_; _↔_; mk⇔; mk↔′) open import Function.Definitions.Core2 using (Surjective) open import Relation.Binary as B hiding (Decidable; _⇔_) open import Relation.Binary.PropositionalEquality as P - using (_≡_; _≢_; refl; sym; trans; cong; subst; _≗_; module ≡-Reasoning) + using (_≡_; _≢_; refl; sym; trans; cong; cong₂; subst; _≗_; module ≡-Reasoning) open import Relation.Nullary.Decidable as Dec using (map′) open import Relation.Nullary.Reflects open import Relation.Nullary.Negation using (contradiction) @@ -151,6 +153,22 @@ toℕ≤pred[n] (suc {n = suc n} i) = s≤s (toℕ≤pred[n] i) toℕ≤pred[n]′ : ∀ {n} (i : Fin n) → toℕ i ℕ.≤ ℕ.pred n toℕ≤pred[n]′ i = ℕₚ.<⇒≤pred (toℕ _ _ i>k = contradiction (sym combine[i,j]≡combine[k,l]) (<⇒≢ (lemma₁ k l i j i>k)) + + lemma₃ : ∀ {n m} (i : Fin n) (j : Fin m) (k : Fin n) (l : Fin m) → combine i j ≡ combine k l → j ≡ l + lemma₃ i j k l combine[i,j]≡combine[k,l] = combine-injectiveʳ i j l (begin + combine i j ≡⟨ combine[i,j]≡combine[k,l] ⟩ + combine k l ≡˘⟨ cong (λ h → combine h l) (lemma₂ i j k l combine[i,j]≡combine[k,l]) ⟩ + combine i l ∎) + where open ≡-Reasoning + +combine-surjective : ∀ {n m} (i : Fin (n ℕ.* m)) → Σ[ j ∈ Fin n ] Σ[ k ∈ Fin m ] combine j k ≡ i +combine-surjective {n} {m} i with remQuot {n} m i | P.inspect (remQuot {n} m) i +... | j , k | P.[ eq ] = j , k , (begin + combine j k ≡˘⟨ uncurry (cong₂ combine) (,-injective eq) ⟩ + uncurry combine (remQuot {n} m i) ≡⟨ combine-remQuot {n} m i ⟩ + i ∎) + where open ≡-Reasoning + ------------------------------------------------------------------------ -- Bundles