diff options
author | Helmut Grohne <helmut@subdivi.de> | 2012-04-19 11:52:27 +0200 |
---|---|---|
committer | Helmut Grohne <helmut@subdivi.de> | 2012-04-19 11:52:27 +0200 |
commit | 7c3e2c61e55aa876f88fbd34c94ccfb0a8c715d4 (patch) | |
tree | da62a6b97b8744c0bb7b9067b9007a8960004246 /FinMap.agda | |
parent | dabd8455638d0cea486dd94dcdd13729077018d7 (diff) | |
download | bidiragda-7c3e2c61e55aa876f88fbd34c94ccfb0a8c715d4.tar.gz |
move lemma-just!=nothing to FinMap and use it there
Diffstat (limited to 'FinMap.agda')
-rw-r--r-- | FinMap.agda | 14 |
1 files changed, 4 insertions, 10 deletions
diff --git a/FinMap.agda b/FinMap.agda index 03a304b..81862f6 100644 --- a/FinMap.agda +++ b/FinMap.agda @@ -12,7 +12,7 @@ open import Function using (id ; _∘_ ; flip) open import Relation.Nullary using (¬_ ; yes ; no) open import Relation.Nullary.Negation using (contradiction ; contraposition) open import Relation.Binary.Core using (_≡_ ; refl) -open import Relation.Binary.PropositionalEquality using (cong ; sym ; _≗_) +open import Relation.Binary.PropositionalEquality using (cong ; sym ; _≗_ ; trans) open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎) FinMapMaybe : ℕ → Set → Set @@ -46,6 +46,8 @@ union m1 m2 = tabulate (λ f → maybe′ id (lookup f m2) (lookupM f m1)) restrict : {A : Set} {n : ℕ} → (Fin n → A) → List (Fin n) → FinMapMaybe n A restrict f is = fromAscList (zip is (map f is)) +lemma-just≢nothing : {A Whatever : Set} {a : A} → _≡_ {_} {Maybe A} (just a) nothing → Whatever +lemma-just≢nothing () lemma-insert-same : {τ : Set} {n : ℕ} → (m : FinMapMaybe n τ) → (f : Fin n) → (a : τ) → lookupM f m ≡ just a → m ≡ insert f a m lemma-insert-same [] () a p @@ -70,15 +72,7 @@ lemma-from-just : {A : Set} → {x y : A} → _≡_ {_} {Maybe A} (just x) (just lemma-from-just refl = refl lemma-lookupM-restrict : {A : Set} {n : ℕ} → (i : Fin n) → (f : Fin n → A) → (is : List (Fin n)) → (a : A) → lookupM i (restrict f is) ≡ just a → f i ≡ a -lemma-lookupM-restrict {A} i f [] a p with begin - just a - ≡⟨ sym p ⟩ - lookupM i (restrict f []) - ≡⟨ refl ⟩ - lookupM i empty - ≡⟨ lemma-lookupM-empty i ⟩ - nothing ∎ -lemma-lookupM-restrict i f [] a p | () +lemma-lookupM-restrict {A} i f [] a p = lemma-just≢nothing (trans (sym p) (lemma-lookupM-empty i)) lemma-lookupM-restrict i f (i' ∷ is) a p with i ≟ i' lemma-lookupM-restrict i f (.i ∷ is) a p | yes refl = lemma-from-just (begin just (f i) |