Simplify proofs about 'loop' using concatenation lemma

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
2024-04-29 21:28:21 -07:00
parent c574ca9c56
commit 91b5d108f6
2 changed files with 33 additions and 6 deletions

View File

@@ -3,7 +3,7 @@ module Utils where
open import Agda.Primitive using () renaming (_⊔_ to _⊔_)
open import Data.Product as Prod using ()
open import Data.Nat using (; suc)
open import Data.List using (List; cartesianProduct; []; _∷_; _++_) renaming (map to mapˡ)
open import Data.List using (List; cartesianProduct; []; _∷_; _++_; foldr) renaming (map to mapˡ)
open import Data.List.Membership.Propositional using (_∈_)
open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
open import Data.List.Relation.Unary.All using (All; []; _∷_; map)
@@ -86,3 +86,8 @@ proj₂ (_ , v) = v
x xs y ys (x Prod., y) cartesianProduct xs ys
∈-cartesianProduct {x = x} (here refl) y∈ys = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (x Prod.,_) y∈ys)
∈-cartesianProduct {x = x} {xs = x' _} {ys = ys} (there x∈rest) y∈ys = ListMemProp.∈-++⁺ʳ (mapˡ (x' Prod.,_) ys) (∈-cartesianProduct x∈rest y∈ys)
concat-∈ : {a} {A : Set a} {x : A} {l : List A} {ls : List (List A)}
x l l ls x foldr _++_ [] ls
concat-∈ x∈l (here refl) = ListMemProp.∈-++⁺ˡ x∈l
concat-∈ {ls = l' ls'} x∈l (there l∈ls') = ListMemProp.∈-++⁺ʳ l' (concat-∈ x∈l l∈ls')