@@ -20,7 +20,7 @@ open import Function.Base using (flip; id; _∘_)
2020open import Function.Bundles using (_↔_; mk↔ₛ′)
2121open import Level using (Level)
2222open import Relation.Binary.PropositionalEquality.Core
23- using (_≡_; refl; cong; resp)
23+ using (_≡_; refl; trans; cong; resp)
2424open import Relation.Unary using (Pred; _⊆_)
2525
2626private
@@ -42,7 +42,8 @@ private
4242
4343find-∈ : (x∈xs : x ∈ xs) → find x∈xs ≡ (x , x∈xs , refl)
4444find-∈ (here refl) = refl
45- find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl
45+ find-∈ (there x∈xs)
46+ = cong (λ where (x , x∈xs , eq) → x , there x∈xs , eq) (find-∈ x∈xs)
4647
4748------------------------------------------------------------------------
4849-- Lemmas relating map and find.
@@ -59,22 +60,24 @@ module _ {P : Pred A p} where
5960 let x , x∈xs , px = find p in
6061 find (Any.map f p) ≡ (x , x∈xs , f px)
6162 find∘map (here p) f = refl
62- find∘map (there p) f rewrite find∘map p f = refl
63+ find∘map (there p) f
64+ = cong (λ where (x , x∈xs , eq) → x , there x∈xs , eq) (find∘map p f)
6365
6466------------------------------------------------------------------------
6567-- Any can be expressed using _∈_
6668
6769module _ {P : Pred A p} where
6870
6971 ∃∈-Any : (∃ λ x → x ∈ xs × P x) → Any P xs
70- ∃∈-Any (x , x∈xs , px) = lose {P = P} x∈xs px
72+ ∃∈-Any (x , x∈xs , px) = lose x∈xs px
7173
7274 ∃∈-Any∘find : (p : Any P xs) → ∃∈-Any (find p) ≡ p
7375 ∃∈-Any∘find p = map∘find p refl
7476
7577 find∘∃∈-Any : (p : ∃ λ x → x ∈ xs × P x) → find (∃∈-Any p) ≡ p
7678 find∘∃∈-Any p@(x , x∈xs , px)
77- rewrite find∘map x∈xs (flip (resp P) px) | find-∈ x∈xs = refl
79+ = trans (find∘map x∈xs (flip (resp P) px))
80+ (cong (λ (x , x∈xs , eq) → x , x∈xs , resp P eq px) (find-∈ x∈xs))
7881
7982 Any↔ : (∃ λ x → x ∈ xs × P x) ↔ Any P xs
8083 Any↔ = mk↔ₛ′ ∃∈-Any find ∃∈-Any∘find find∘∃∈-Any
0 commit comments