summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--Bidir.agda12
-rw-r--r--CheckInsert.agda10
2 files changed, 11 insertions, 11 deletions
diff --git a/Bidir.agda b/Bidir.agda
index 98cb93c..721b2b2 100644
--- a/Bidir.agda
+++ b/Bidir.agda
@@ -16,7 +16,7 @@ open import Function using (id ; _∘_ ; flip)
open import Relation.Nullary using (yes ; no ; ¬_)
open import Relation.Nullary.Negation using (contradiction)
open import Relation.Binary.Core using (_≡_ ; refl)
-open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; Reveal_is_ ; _≗_ ; trans)
+open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; _≗_ ; trans)
open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
open import FinMap
@@ -66,8 +66,8 @@ lemma-∉-lookupM-assoc eq i [] [] h ph i∉is = begin
≡⟨ lemma-lookupM-empty i ⟩
nothing ∎
lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h ph i∉is with assoc eq is' xs' | inspect (assoc eq is') xs'
-lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h () i∉is | nothing | Reveal_is_.[_] ph'
-lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h ph i∉is | just h' | Reveal_is_.[_] ph' = apply-checkInsertProof eq i' x' h' record {
+lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h () i∉is | nothing | [ ph' ]
+lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h ph i∉is | just h' | [ ph' ] = apply-checkInsertProof eq i' x' h' record {
same = λ lookupM-i'-h'≡just-x' → begin
lookupM i h
≡⟨ cong (lookupM i) (lemma-from-just (trans (sym ph) (lemma-checkInsert-same eq i' x' h' lookupM-i'-h'≡just-x'))) ⟩
@@ -91,8 +91,8 @@ _in-domain-of_ is h = All (λ i → ∃ λ x → lookupM i h ≡ just x) is
lemma-assoc-domain : {m n : ℕ} {A : Set} → (eq : EqInst A) → (is : Vec (Fin n) m) → (xs : Vec A m) → (h : FinMapMaybe n A) → assoc eq is xs ≡ just h → (toList is) in-domain-of h
lemma-assoc-domain eq [] [] h ph = Data.List.All.[]
lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h ph with assoc eq is' xs' | inspect (assoc eq is') xs'
-lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h () | nothing | ph'
-lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h ph | just h' | Reveal_is_.[_] ph' = apply-checkInsertProof eq i' x' h' record {
+lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h () | nothing | [ ph' ]
+lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h ph | just h' | [ ph' ] = apply-checkInsertProof eq i' x' h' record {
same = λ lookupM-i'-h'≡just-x' → Data.List.All._∷_
(x' , (trans (cong (lookupM i') (lemma-from-just (trans (sym ph) (lemma-checkInsert-same eq i' x' h' lookupM-i'-h'≡just-x')))) lookupM-i'-h'≡just-x'))
(lemma-assoc-domain eq is' xs' h (trans ph' (trans (sym (lemma-checkInsert-same eq i' x' h' lookupM-i'-h'≡just-x')) ph)))
@@ -132,7 +132,7 @@ lemma-2 : {Ï„ : Set} {m n : â„•} → (eq : EqInst Ï„) → (is : Vec (Fin n) m) â
lemma-2 eq [] [] h p = refl
lemma-2 eq (i ∷ is) (x ∷ xs) h p with assoc eq is xs | inspect (assoc eq is) xs
lemma-2 eq (i ∷ is) (x ∷ xs) h () | nothing | _
-lemma-2 eq (i ∷ is) (x ∷ xs) h p | just h' | Reveal_is_.[_] ir = begin
+lemma-2 eq (i ∷ is) (x ∷ xs) h p | just h' | [ ir ] = begin
map (flip lookupM h) (i ∷ is)
≡⟨ refl ⟩
lookupM i h ∷ map (flip lookupM h) is
diff --git a/CheckInsert.agda b/CheckInsert.agda
index 408a5b2..c482423 100644
--- a/CheckInsert.agda
+++ b/CheckInsert.agda
@@ -8,7 +8,7 @@ open import Data.List using (List ; [] ; _∷_)
open import Relation.Nullary using (Dec ; yes ; no ; ¬_)
open import Relation.Nullary.Negation using (contradiction)
open import Relation.Binary.Core using (_≡_ ; refl)
-open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; Reveal_is_ ; trans)
+open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; trans)
open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
open import FinMap
@@ -32,9 +32,9 @@ record checkInsertProof {A : Set} {n : â„•} (eq : EqInst A) (i : Fin n) (x : A)
apply-checkInsertProof : {A P : Set} {n : ℕ} → (eq : EqInst A) → (i : Fin n) → (x : A) → (m : FinMapMaybe n A) → checkInsertProof eq i x m P → P
apply-checkInsertProof eq i x m rp with lookupM i m | inspect (lookupM i) m
apply-checkInsertProof eq i x m rp | just x' | il with eq x x'
-apply-checkInsertProof eq i x m rp | just .x | Reveal_is_.[_] il | yes refl = checkInsertProof.same rp il
-apply-checkInsertProof eq i x m rp | just x' | Reveal_is_.[_] il | no x≢x' = checkInsertProof.wrong rp x' x≢x' il
-apply-checkInsertProof eq i x m rp | nothing | Reveal_is_.[_] il = checkInsertProof.new rp il
+apply-checkInsertProof eq i x m rp | just .x | [ il ] | yes refl = checkInsertProof.same rp il
+apply-checkInsertProof eq i x m rp | just x' | [ il ] | no x≢x' = checkInsertProof.wrong rp x' x≢x' il
+apply-checkInsertProof eq i x m rp | nothing | [ il ] = checkInsertProof.new rp il
lemma-checkInsert-same : {A : Set} {n : ℕ} → (eq : EqInst A) → (i : Fin n) → (x : A) → (m : FinMapMaybe n A) → lookupM i m ≡ just x → checkInsert eq i x m ≡ just m
lemma-checkInsert-same eq i x m p with lookupM i m
@@ -75,7 +75,7 @@ lemma-checkInsert-restrict {Ï„} eq f i is = apply-checkInsertProof eq i (f i) (r
lemma-lookupM-checkInsert : {A : Set} {n : ℕ} → (eq : EqInst A) → (i j : Fin n) → (x y : A) → (h h' : FinMapMaybe n A) → lookupM i h ≡ just x → checkInsert eq j y h ≡ just h' → lookupM i h' ≡ just x
lemma-lookupM-checkInsert eq i j x y h h' pl ph' with lookupM j h | inspect (lookupM j) h
lemma-lookupM-checkInsert eq i j x y h .(insert j y h) pl refl | nothing | pl' with i ≟ j
-lemma-lookupM-checkInsert eq i .i x y h .(insert i y h) pl refl | nothing | Reveal_is_.[_] pl' | yes refl = lemma-just≢nothing (begin just x ≡⟨ sym pl ⟩ lookupM i h ≡⟨ pl' ⟩ (nothing ∎))
+lemma-lookupM-checkInsert eq i .i x y h .(insert i y h) pl refl | nothing | [ pl' ] | yes refl = lemma-just≢nothing (begin just x ≡⟨ sym pl ⟩ lookupM i h ≡⟨ pl' ⟩ (nothing ∎))
lemma-lookupM-checkInsert eq i j x y h .(insert j y h) pl refl | nothing | pl' | no ¬p = begin
lookupM i (insert j y h)
≡⟨ sym (lemma-lookupM-insert-other i j y h ¬p) ⟩