From 066861f9cdde4ded6c5442508bef1a27576c68d7 Mon Sep 17 00:00:00 2001 From: Helmut Grohne Date: Tue, 17 Dec 2013 08:22:12 +0100 Subject: update bff implementation to use delete In the presence of shape-changing updates, bff needs to shrink one of the mappings before unifying them. As long the shape does not change, the union becomes a disjoint union. With this insight we can adapt the proof of theorem-1 using the adapted lemma-disjoint-union. Unfortunately theorem-2 requires more work and assoc-enough becomes non-trivial due to the introduction of mapMV. --- Precond.agda | 2 ++ 1 file changed, 2 insertions(+) (limited to 'Precond.agda') diff --git a/Precond.agda b/Precond.agda index e4699dc..16e452b 100644 --- a/Precond.agda +++ b/Precond.agda @@ -28,11 +28,13 @@ import Bidir open BFF.VecBFF Carrier deq using (get-type ; assoc ; enumerate ; denumerate ; bff) +{- assoc-enough : {getlen : ℕ → ℕ} (get : get-type getlen) → {m : ℕ} → (s : Vec Carrier m) → (v : Vec Carrier (getlen m)) → ∃ (λ h → assoc (get (enumerate s)) v ≡ just h) → ∃ λ u → bff get s v ≡ just u assoc-enough get s v (h , p) = u , cong (_<$>_ (flip map s′ ∘ flip lookup) ∘ (_<$>_ (flip union g))) p where s′ = enumerate s g = fromFunc (denumerate s) u = map (flip lookup (union h g)) s′ +-} data All-different {A : Set} : List A → Set where different-[] : All-different [] -- cgit v1.2.3 From 88ac52ceabaf2bf4a3f27293a8e04a3daaf8b907 Mon Sep 17 00:00:00 2001 From: Helmut Grohne Date: Tue, 17 Dec 2013 08:49:24 +0100 Subject: refactor to get rid of FinMap without Maybe entirely The union was the only user of this type and now it uses only partial mappings. So drop remaining uses of FinMap and make everything work with FinMapMaybe instead. --- BFF.agda | 2 +- Bidir.agda | 12 ++++++------ FinMap.agda | 49 ++++++++++++++++++++----------------------------- Precond.agda | 2 +- 4 files changed, 28 insertions(+), 37 deletions(-) (limited to 'Precond.agda') diff --git a/BFF.agda b/BFF.agda index f7ddd30..896350a 100644 --- a/BFF.agda +++ b/BFF.agda @@ -35,7 +35,7 @@ module VecBFF (Carrier : Set) (deq : Decidable {A = Carrier} _≡_) where bff : {getlen : ℕ → ℕ} → (get-type getlen) → ({n : ℕ} → Vec Carrier n → Vec Carrier (getlen n) → Maybe (Vec Carrier n)) bff get s v = let s′ = enumerate s t′ = get s′ - g = partialize (fromFunc (denumerate s)) + g = fromFunc (denumerate s) g′ = delete-many t′ g h = assoc t′ v h′ = (flip union g′) <$> h diff --git a/Bidir.agda b/Bidir.agda index c74601a..1e65bab 100644 --- a/Bidir.agda +++ b/Bidir.agda @@ -113,20 +113,20 @@ theorem-1 get s = begin ≡⟨ cong (h′↦r ∘ h↦h′) (lemma-1 (denumerate s) (get (enumerate s))) ⟩ (h′↦r ∘ h↦h′ ∘ just) (restrict (denumerate s) (toList (get (enumerate s)))) ≡⟨ refl ⟩ - (h′↦r ∘ just) (union (restrict (denumerate s) (toList (get (enumerate s)))) (delete-many (get (enumerate s)) (partialize (fromFunc (denumerate s))))) + (h′↦r ∘ just) (union (restrict (denumerate s) (toList (get (enumerate s)))) (delete-many (get (enumerate s)) (fromFunc (denumerate s)))) ≡⟨ cong (h′↦r ∘ just) (lemma-disjoint-union (denumerate s) (get (enumerate s))) ⟩ - (h′↦r ∘ just) (partialize (fromFunc (denumerate s))) + (h′↦r ∘ just) (fromFunc (denumerate s)) ≡⟨ refl ⟩ - mapMV (flip lookupVec (partialize (fromFunc (denumerate s)))) (enumerate s) - ≡⟨ cong (flip mapMV (enumerate s) ∘ flip lookupVec) (lemma-partialize-fromFunc (denumerate s)) ⟩ - mapMV (flip lookupVec (fromFunc (Maybe.just ∘ denumerate s))) (enumerate s) + mapMV (flip lookupVec (fromFunc (denumerate s))) (enumerate s) + ≡⟨ cong (flip mapMV (enumerate s) ∘ flip lookupVec) (lemma-fromFunc-tabulate (denumerate s)) ⟩ + mapMV (flip lookupVec (tabulate (Maybe.just ∘ denumerate s))) (enumerate s) ≡⟨ mapMV-cong (lookup∘tabulate (Maybe.just ∘ denumerate s)) (enumerate s) ⟩ mapMV (Maybe.just ∘ denumerate s) (enumerate s) ≡⟨ mapMV-purity (denumerate s) (enumerate s) ⟩ just (map (denumerate s) (enumerate s)) ≡⟨ cong just (lemma-map-denumerate-enumerate s) ⟩ just s ∎ - where h↦h′ = _<$>_ (flip union (delete-many (get (enumerate s)) (partialize (fromFunc (denumerate s))))) + where h↦h′ = _<$>_ (flip union (delete-many (get (enumerate s)) (fromFunc (denumerate s)))) h′↦r = flip _>>=_ (flip mapMV (enumerate s) ∘ flip lookupVec) diff --git a/FinMap.agda b/FinMap.agda index 8cde5a6..c125c47 100644 --- a/FinMap.agda +++ b/FinMap.agda @@ -33,17 +33,11 @@ fromAscList : {A : Set} {n : ℕ} → List (Fin n × A) → FinMapMaybe n A fromAscList [] = empty fromAscList ((f , a) ∷ xs) = insert f a (fromAscList xs) -FinMap : ℕ → Set → Set -FinMap n A = Vec A n - -lookup : {A : Set} {n : ℕ} → Fin n → FinMap n A → A -lookup = lookupVec - -fromFunc : {A : Set} {n : ℕ} → (Fin n → A) → FinMap n A -fromFunc = tabulate +fromFunc : {A : Set} {n : ℕ} → (Fin n → A) → FinMapMaybe n A +fromFunc = mapV just ∘ tabulate union : {A : Set} {n : ℕ} → FinMapMaybe n A → FinMapMaybe n A → FinMapMaybe n A -union m1 m2 = fromFunc (λ f → maybe′ just (lookupM f m2) (lookupM f m1)) +union m1 m2 = tabulate (λ f → maybe′ just (lookupM 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)) @@ -54,9 +48,6 @@ delete i m = m [ i ]≔ nothing delete-many : {A : Set} {n m : ℕ} → Vec (Fin n) m → FinMapMaybe n A → FinMapMaybe n A delete-many = flip (foldr (const _) delete) -partialize : {A : Set} {n : ℕ} → FinMap n A → FinMapMaybe n A -partialize = mapV just - lemma-just≢nothing : {A Whatever : Set} {a : A} {ma : Maybe A} → ma ≡ just a → ma ≡ nothing → Whatever lemma-just≢nothing refl () @@ -99,9 +90,9 @@ lemma-tabulate-∘ : {n : ℕ} {A : Set} → {f g : Fin n → A} → f ≗ g → lemma-tabulate-∘ {zero} {_} {f} {g} f≗g = refl lemma-tabulate-∘ {suc n} {_} {f} {g} f≗g = cong₂ _∷_ (f≗g zero) (lemma-tabulate-∘ (f≗g ∘ suc)) -lemma-partialize-fromFunc : {n : ℕ} {A : Set} → (f : Fin n → A) → partialize (fromFunc f) ≡ fromFunc (just ∘ f) -lemma-partialize-fromFunc {zero} f = refl -lemma-partialize-fromFunc {suc _} f = cong (_∷_ (just (f zero))) (lemma-partialize-fromFunc (f ∘ suc)) +lemma-fromFunc-tabulate : {n : ℕ} {A : Set} → (f : Fin n → A) → fromFunc f ≡ tabulate (just ∘ f) +lemma-fromFunc-tabulate {zero} f = refl +lemma-fromFunc-tabulate {suc _} f = cong (_∷_ (just (f zero))) (lemma-fromFunc-tabulate (f ∘ suc)) lemma-lookupM-delete : {n : ℕ} {A : Set} {i j : Fin n} → (f : FinMapMaybe n A) → i ≢ j → lookupM i (delete j f) ≡ lookupM i f lemma-lookupM-delete {i = zero} {j = zero} (_ ∷ _) p with p refl @@ -110,24 +101,24 @@ lemma-lookupM-delete {i = zero} {j = suc j} (_ ∷ _) p = refl lemma-lookupM-delete {i = suc i} {j = zero} (x ∷ xs) p = refl lemma-lookupM-delete {i = suc i} {j = suc j} (x ∷ xs) p = lemma-lookupM-delete xs (p ∘ cong suc) -lemma-disjoint-union : {n m : ℕ} {A : Set} → (f : Fin n → A) → (t : Vec (Fin n) m) → union (restrict f (toList t)) (delete-many t (partialize (fromFunc f))) ≡ partialize (fromFunc f) -lemma-disjoint-union {n} {m} f t = trans (lemma-tabulate-∘ (lemma-inner t)) (sym (lemma-partialize-fromFunc f)) - where lemma-inner : {m : ℕ} → (t : Vec (Fin n) m) → (x : Fin n) → maybe′ just (lookupM x (delete-many t (partialize (fromFunc f)))) (lookupM x (restrict f (toList t))) ≡ just (f x) +lemma-disjoint-union : {n m : ℕ} {A : Set} → (f : Fin n → A) → (t : Vec (Fin n) m) → union (restrict f (toList t)) (delete-many t (fromFunc f)) ≡ fromFunc f +lemma-disjoint-union {n} {m} f t = trans (lemma-tabulate-∘ (lemma-inner t)) (sym (lemma-fromFunc-tabulate f)) + where lemma-inner : {m : ℕ} → (t : Vec (Fin n) m) → (x : Fin n) → maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f (toList t))) ≡ just (f x) lemma-inner [] x = begin - maybe′ just (lookupM x (partialize (fromFunc f))) (lookupM x empty) - ≡⟨ cong (maybe′ just (lookupM x (partialize (fromFunc f)))) (lemma-lookupM-empty x) ⟩ - lookupM x (partialize (fromFunc f)) - ≡⟨ cong (lookupM x) (lemma-partialize-fromFunc f) ⟩ - lookupM x (fromFunc (just ∘ f)) + maybe′ just (lookupM x (fromFunc f)) (lookupM x empty) + ≡⟨ cong (maybe′ just (lookupM x (fromFunc f))) (lemma-lookupM-empty x) ⟩ + lookupM x (fromFunc f) + ≡⟨ cong (lookupM x) (lemma-fromFunc-tabulate f) ⟩ + lookupM x (tabulate (just ∘ f)) ≡⟨ lookup∘tabulate (just ∘ f) x ⟩ just (f x) ∎ lemma-inner (t ∷ ts) x with x ≟ t - lemma-inner (.x ∷ ts) x | yes refl = cong (maybe′ just (lookupM x (delete-many (x ∷ ts) (partialize (fromFunc f))))) (lemma-lookupM-insert x (f x) (restrict f (toList ts))) + lemma-inner (.x ∷ ts) x | yes refl = cong (maybe′ just (lookupM x (delete-many (x ∷ ts) (fromFunc f)))) (lemma-lookupM-insert x (f x) (restrict f (toList ts))) lemma-inner (t ∷ ts) x | no ¬p = begin - maybe′ just (lookupM x (delete-many (t ∷ ts) (partialize (fromFunc f)))) (lookupM x (restrict f (toList (t ∷ ts)))) - ≡⟨ cong (maybe′ just (lookupM x (delete-many (t ∷ ts) (partialize (fromFunc f))))) (sym (lemma-lookupM-insert-other x t (f t) (restrict f (toList ts)) ¬p)) ⟩ - maybe′ just (lookupM x (delete-many (t ∷ ts) (partialize (fromFunc f)))) (lookupM x (restrict f (toList ts))) - ≡⟨ cong (flip (maybe′ just) (lookupM x (restrict f (toList ts)))) (lemma-lookupM-delete (delete-many ts (partialize (fromFunc f))) ¬p) ⟩ - maybe′ just (lookupM x (delete-many ts (partialize (fromFunc f)))) (lookupM x (restrict f (toList ts))) + maybe′ just (lookupM x (delete-many (t ∷ ts) (fromFunc f))) (lookupM x (restrict f (toList (t ∷ ts)))) + ≡⟨ cong (maybe′ just (lookupM x (delete-many (t ∷ ts) (fromFunc f)))) (sym (lemma-lookupM-insert-other x t (f t) (restrict f (toList ts)) ¬p)) ⟩ + maybe′ just (lookupM x (delete-many (t ∷ ts) (fromFunc f))) (lookupM x (restrict f (toList ts))) + ≡⟨ cong (flip (maybe′ just) (lookupM x (restrict f (toList ts)))) (lemma-lookupM-delete (delete-many ts (fromFunc f)) ¬p) ⟩ + maybe′ just (lookupM x (delete-many ts (fromFunc f))) (lookupM x (restrict f (toList ts))) ≡⟨ lemma-inner ts x ⟩ just (f x) ∎ diff --git a/Precond.agda b/Precond.agda index 16e452b..40ffd79 100644 --- a/Precond.agda +++ b/Precond.agda @@ -20,7 +20,7 @@ open import Function using (flip ; _∘_) open import Relation.Binary.PropositionalEquality using (refl ; cong ; inspect ; [_] ; sym) open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎) -open import FinMap using (FinMap ; FinMapMaybe ; lookupM ; union ; fromFunc ; empty ; insert ; lemma-lookupM-empty) +open import FinMap using (FinMapMaybe ; lookupM ; union ; fromFunc ; empty ; insert ; lemma-lookupM-empty) import CheckInsert open CheckInsert Carrier deq using (checkInsert ; lemma-checkInsert-new ; lemma-lookupM-checkInsert-other) import BFF -- cgit v1.2.3 From d2521627834713a651be0ac22aab0a1cd78df920 Mon Sep 17 00:00:00 2001 From: Helmut Grohne Date: Mon, 27 Jan 2014 09:02:45 +0100 Subject: prove assoc-enough in the presence of delete The old definition of bff had a single failure mode - assoc - and its proof was a single cong. Now we need to show that the other failure mode - mapM (flip lookupM ...) - is eliminated by the success of the former resulting in a larger proof. --- Precond.agda | 84 ++++++++++++++++++++++++++++++++++++++++++++++++++---------- 1 file changed, 71 insertions(+), 13 deletions(-) (limited to 'Precond.agda') diff --git a/Precond.agda b/Precond.agda index 40ffd79..90027b4 100644 --- a/Precond.agda +++ b/Precond.agda @@ -3,38 +3,96 @@ open import Relation.Binary.Core using (Decidable ; _≡_) module Precond (Carrier : Set) (deq : Decidable {A = Carrier} _≡_) where open import Data.Nat using (ℕ) -open import Data.Fin using (Fin) +open import Data.Fin using (Fin ; zero ; suc) +open import Data.Fin.Props using (_≟_) open import Data.List using (List ; [] ; _∷_) import Level import Category.Monad import Category.Functor -open import Data.Maybe using (nothing ; just) +open import Data.Maybe using (Maybe ; nothing ; just ; maybe′) open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_) open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_) -open import Data.Vec using (Vec ; [] ; _∷_ ; map ; lookup ; toList) +open import Data.Vec using (Vec ; [] ; _∷_ ; map ; lookup ; toList ; tabulate) +open import Data.Vec.Properties using (map-cong ; map-∘ ; tabulate-∘) +import Data.List.All open import Data.List.Any using (here ; there) open Data.List.Any.Membership-≡ using (_∉_) open import Data.Maybe using (just) -open import Data.Product using (∃ ; _,_) -open import Function using (flip ; _∘_) +open import Data.Product using (∃ ; _,_ ; proj₁ ; proj₂) +open import Function using (flip ; _∘_ ; id) open import Relation.Binary.PropositionalEquality using (refl ; cong ; inspect ; [_] ; sym) open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎) +open import Relation.Nullary using (yes ; no) -open import FinMap using (FinMapMaybe ; lookupM ; union ; fromFunc ; empty ; insert ; lemma-lookupM-empty) +open import Generic using (mapMV ; sequenceV ; sequence-map) +open import FinMap using (FinMapMaybe ; lookupM ; union ; fromFunc ; empty ; insert ; lemma-lookupM-empty ; delete-many ; lemma-tabulate-∘ ; delete ; lemma-lookupM-delete) import CheckInsert open CheckInsert Carrier deq using (checkInsert ; lemma-checkInsert-new ; lemma-lookupM-checkInsert-other) import BFF -import Bidir +open import Bidir Carrier deq using (_in-domain-of_ ; lemma-assoc-domain ; lemma-just-sequence) open BFF.VecBFF Carrier deq using (get-type ; assoc ; enumerate ; denumerate ; bff) -{- +lemma-lookup-map-just : {n : ℕ} (f : Fin n) {A : Set} (v : Vec A n) → lookup f (map Maybe.just v) ≡ Maybe.just (lookup f v) +lemma-lookup-map-just zero (x ∷ xs) = refl +lemma-lookup-map-just (suc f) (x ∷ xs) = lemma-lookup-map-just f xs + +fromMaybe : {A : Set} → A → Maybe A → A +fromMaybe a ma = maybe′ id a ma + +lemma-maybe-just : {A : Set} → (a : A) → (ma : Maybe A) → maybe′ Maybe.just (just a) ma ≡ Maybe.just (maybe′ id a ma) +lemma-maybe-just a (just x) = refl +lemma-maybe-just a nothing = refl + +lemma-union-delete-fromFunc : {m n : ℕ} {A : Set} {is : Vec (Fin n) m} {h : FinMapMaybe n A} {g : Vec A n} → (toList is) in-domain-of h → ∃ λ v → union h (delete-many is (map just g)) ≡ map just v +lemma-union-delete-fromFunc {is = []} {h = h} {g = g} p = _ , (begin + union h (map just g) + ≡⟨ lemma-tabulate-∘ (λ f → begin + maybe′ just (lookup f (map just g)) (lookup f h) + ≡⟨ cong (flip (maybe′ just) (lookup f h)) (lemma-lookup-map-just f g) ⟩ + maybe′ just (just (lookup f g)) (lookup f h) + ≡⟨ lemma-maybe-just (lookup f g) (lookup f h) ⟩ + just (fromMaybe (lookup f g) (lookup f h)) ∎) ⟩ + tabulate (λ f → just (fromMaybe (lookup f g) (lookup f h))) + ≡⟨ tabulate-∘ just (λ f → fromMaybe (lookup f g) (lookup f h)) ⟩ + map just (tabulate (λ f → fromMaybe (lookup f g) (lookup f h))) ∎) +lemma-union-delete-fromFunc {n = n} {is = i ∷ is} {h = h} {g = g} (p Data.List.All.∷ ps) = _ , (begin + union h (delete i (delete-many is (map just g))) + ≡⟨ lemma-tabulate-∘ (λ f → (begin + maybe′ just (lookupM f (delete i (delete-many is (map just g)))) (lookup f h) + ≡⟨ inner f ⟩ + maybe′ just (lookupM f (delete-many is (map just g))) (lookup f h) ∎)) ⟩ + union h (delete-many is (map just g)) + ≡⟨ proj₂ (lemma-union-delete-fromFunc ps) ⟩ + map just _ ∎) + where inner : (f : Fin n) → maybe′ just (lookupM f (delete i (delete-many is (map just g)))) (lookup f h) ≡ maybe′ just (lookupM f (delete-many is (map just g))) (lookup f h) + inner f with f ≟ i + inner .i | yes refl = begin + maybe′ just (lookupM i (delete i (delete-many is (map just g)))) (lookup i h) + ≡⟨ cong (maybe′ just _) (proj₂ p) ⟩ + just (proj₁ p) + ≡⟨ cong (maybe′ just _) (sym (proj₂ p)) ⟩ + maybe′ just (lookupM i (delete-many is (map just g))) (lookup i h) ∎ + inner f | no f≢i = cong (flip (maybe′ just) (lookup f h)) (lemma-lookupM-delete (delete-many is (map just g)) f≢i) + assoc-enough : {getlen : ℕ → ℕ} (get : get-type getlen) → {m : ℕ} → (s : Vec Carrier m) → (v : Vec Carrier (getlen m)) → ∃ (λ h → assoc (get (enumerate s)) v ≡ just h) → ∃ λ u → bff get s v ≡ just u -assoc-enough get s v (h , p) = u , cong (_<$>_ (flip map s′ ∘ flip lookup) ∘ (_<$>_ (flip union g))) p - where s′ = enumerate s - g = fromFunc (denumerate s) - u = map (flip lookup (union h g)) s′ --} +assoc-enough get s v (h , p) = let w , pw = lemma-union-delete-fromFunc (lemma-assoc-domain (get s′) v h p) in _ , (begin + bff get s v + ≡⟨ cong (flip _>>=_ (flip mapMV s′ ∘ flip lookupM) ∘ _<$>_ (flip union g′)) p ⟩ + mapMV (flip lookupM (union h g′)) s′ + ≡⟨ sym (sequence-map (flip lookupM (union h g′)) s′) ⟩ + sequenceV (map (flip lookupM (union h g′)) s′) + ≡⟨ cong (sequenceV ∘ flip map s′ ∘ flip lookupM) pw ⟩ + sequenceV (map (flip lookupM (map just w)) s′) + ≡⟨ cong sequenceV (map-cong (λ f → lemma-lookup-map-just f w) s′) ⟩ + sequenceV (map (Maybe.just ∘ flip lookup w) s′) + ≡⟨ cong sequenceV (map-∘ just (flip lookup w) s′) ⟩ + sequenceV (map Maybe.just (map (flip lookup w) s′)) + ≡⟨ lemma-just-sequence (map (flip lookup w) s′) ⟩ + just (map (flip lookup w) s′) ∎) + where s′ = enumerate s + g = fromFunc (denumerate s) + g′ = delete-many (get s′) g data All-different {A : Set} : List A → Set where different-[] : All-different [] -- cgit v1.2.3 From 71025b5f1d0a11b0cf373192210b293a77d45c04 Mon Sep 17 00:00:00 2001 From: Helmut Grohne Date: Mon, 27 Jan 2014 09:31:56 +0100 Subject: cleanup unused functions and useless steps --- Bidir.agda | 10 +--------- Generic.agda | 10 ---------- Precond.agda | 26 ++++++++++---------------- 3 files changed, 11 insertions(+), 35 deletions(-) (limited to 'Precond.agda') diff --git a/Bidir.agda b/Bidir.agda index eef8bff..05ee066 100644 --- a/Bidir.agda +++ b/Bidir.agda @@ -22,7 +22,7 @@ open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨ import FreeTheorems open FreeTheorems.VecVec using (get-type ; free-theorem) -open import Generic using (just-injective ; map-just-injective ; mapMV ; mapMV-cong ; mapMV-purity ; sequenceV ; sequence-map) +open import Generic using (just-injective ; mapMV ; mapMV-cong ; mapMV-purity ; sequenceV ; sequence-map) open import FinMap import CheckInsert open CheckInsert Carrier deq @@ -150,14 +150,6 @@ lemma->>=-just : {A B : Set} (ma : Maybe A) {f : A → Maybe B} {b : B} → (ma lemma->>=-just (just a) p = a , refl lemma->>=-just nothing () -lemma-mapMV-just : {A B : Set} {n : ℕ} {f : A → Maybe B} {s : Vec A n} {v : Vec B n} → mapMV f s ≡ just v → All (λ x → ∃ λ y → f x ≡ just y) (toList s) -lemma-mapMV-just {s = []} p = Data.List.All.[] -lemma-mapMV-just {f = f} {s = x ∷ xs} p with f x | inspect f x -lemma-mapMV-just {s = x ∷ xs} () | nothing | _ -lemma-mapMV-just {f = f} {s = x ∷ xs} p | just y | [ py ] with mapMV f xs | inspect (mapMV f) xs -lemma-mapMV-just {s = x ∷ xs} () | just y | [ py ] | nothing | _ -lemma-mapMV-just {s = x ∷ xs} p | just y | [ py ] | just ys | [ pys ] = (y , py) Data.List.All.∷ (lemma-mapMV-just pys) - lemma-just-sequence : {A : Set} {n : ℕ} → (v : Vec A n) → sequenceV (map just v) ≡ just v lemma-just-sequence [] = refl lemma-just-sequence (x ∷ xs) rewrite lemma-just-sequence xs = refl diff --git a/Generic.agda b/Generic.agda index 2ef374c..4e55c21 100644 --- a/Generic.agda +++ b/Generic.agda @@ -15,10 +15,6 @@ open import Relation.Binary.PropositionalEquality using (_≗_ ; cong ; subst ; open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_) open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_) -∷-injective : {A : Set} {n : ℕ} {x y : A} {xs ys : Vec A n} → - (x ∷V xs) ≡ (y ∷V ys) → x ≡ y × xs ≡ ys -∷-injective refl = refl , refl - just-injective : {A : Set} → {x y : A} → Maybe.just x ≡ Maybe.just y → x ≡ y just-injective refl = refl @@ -26,12 +22,6 @@ length-replicate : {A : Set} {a : A} → (n : ℕ) → length (replicate n a) length-replicate zero = refl length-replicate (suc n) = cong suc (length-replicate n) -map-just-injective : {A : Set} {n : ℕ} {xs ys : Vec A n} → - map Maybe.just xs ≡ map Maybe.just ys → xs ≡ ys -map-just-injective {xs = []V} {ys = []V} p = refl -map-just-injective {xs = x ∷V xs′} {ys = y ∷V ys′} p with ∷-injective p -map-just-injective {xs = x ∷V xs′} {ys = .x ∷V ys′} p | refl , p′ = cong (_∷V_ x) (map-just-injective p′) - mapMV : {A B : Set} {n : ℕ} → (A → Maybe B) → Vec A n → Maybe (Vec B n) mapMV f []V = just []V mapMV f (x ∷V xs) = (f x) >>= (λ y → (_∷V_ y) <$> (mapMV f xs)) diff --git a/Precond.agda b/Precond.agda index 90027b4..6aba291 100644 --- a/Precond.agda +++ b/Precond.agda @@ -18,7 +18,7 @@ import Data.List.All open import Data.List.Any using (here ; there) open Data.List.Any.Membership-≡ using (_∉_) open import Data.Maybe using (just) -open import Data.Product using (∃ ; _,_ ; proj₁ ; proj₂) +open import Data.Product using (∃ ; _,_ ; proj₂) open import Function using (flip ; _∘_ ; id) open import Relation.Binary.PropositionalEquality using (refl ; cong ; inspect ; [_] ; sym) open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎) @@ -37,9 +37,6 @@ lemma-lookup-map-just : {n : ℕ} (f : Fin n) {A : Set} (v : Vec A n) → lookup lemma-lookup-map-just zero (x ∷ xs) = refl lemma-lookup-map-just (suc f) (x ∷ xs) = lemma-lookup-map-just f xs -fromMaybe : {A : Set} → A → Maybe A → A -fromMaybe a ma = maybe′ id a ma - lemma-maybe-just : {A : Set} → (a : A) → (ma : Maybe A) → maybe′ Maybe.just (just a) ma ≡ Maybe.just (maybe′ id a ma) lemma-maybe-just a (just x) = refl lemma-maybe-just a nothing = refl @@ -52,16 +49,13 @@ lemma-union-delete-fromFunc {is = []} {h = h} {g = g} p = _ , (begin ≡⟨ cong (flip (maybe′ just) (lookup f h)) (lemma-lookup-map-just f g) ⟩ maybe′ just (just (lookup f g)) (lookup f h) ≡⟨ lemma-maybe-just (lookup f g) (lookup f h) ⟩ - just (fromMaybe (lookup f g) (lookup f h)) ∎) ⟩ - tabulate (λ f → just (fromMaybe (lookup f g) (lookup f h))) - ≡⟨ tabulate-∘ just (λ f → fromMaybe (lookup f g) (lookup f h)) ⟩ - map just (tabulate (λ f → fromMaybe (lookup f g) (lookup f h))) ∎) -lemma-union-delete-fromFunc {n = n} {is = i ∷ is} {h = h} {g = g} (p Data.List.All.∷ ps) = _ , (begin + just (maybe′ id (lookup f g) (lookup f h)) ∎) ⟩ + tabulate (λ f → just (maybe′ id (lookup f g) (lookup f h))) + ≡⟨ tabulate-∘ just (λ f → maybe′ id (lookup f g) (lookup f h)) ⟩ + map just (tabulate (λ f → maybe′ id (lookup f g) (lookup f h))) ∎) +lemma-union-delete-fromFunc {n = n} {is = i ∷ is} {h = h} {g = g} ((x , px) Data.List.All.∷ ps) = _ , (begin union h (delete i (delete-many is (map just g))) - ≡⟨ lemma-tabulate-∘ (λ f → (begin - maybe′ just (lookupM f (delete i (delete-many is (map just g)))) (lookup f h) - ≡⟨ inner f ⟩ - maybe′ just (lookupM f (delete-many is (map just g))) (lookup f h) ∎)) ⟩ + ≡⟨ lemma-tabulate-∘ inner ⟩ union h (delete-many is (map just g)) ≡⟨ proj₂ (lemma-union-delete-fromFunc ps) ⟩ map just _ ∎) @@ -69,9 +63,9 @@ lemma-union-delete-fromFunc {n = n} {is = i ∷ is} {h = h} {g = g} (p Data.List inner f with f ≟ i inner .i | yes refl = begin maybe′ just (lookupM i (delete i (delete-many is (map just g)))) (lookup i h) - ≡⟨ cong (maybe′ just _) (proj₂ p) ⟩ - just (proj₁ p) - ≡⟨ cong (maybe′ just _) (sym (proj₂ p)) ⟩ + ≡⟨ cong (maybe′ just _) px ⟩ + just x + ≡⟨ cong (maybe′ just _) (sym px) ⟩ maybe′ just (lookupM i (delete-many is (map just g))) (lookup i h) ∎ inner f | no f≢i = cong (flip (maybe′ just) (lookup f h)) (lemma-lookupM-delete (delete-many is (map just g)) f≢i) -- cgit v1.2.3