@@ -13,8 +13,7 @@ open import Data.Bool.Properties using (T-∨; T-≡)
13
13
open import Data.Empty using (⊥)
14
14
open import Data.Fin.Base using (Fin; zero; suc)
15
15
open import Data.List.Base as List hiding (find)
16
- open import Data.List.Properties using (ʳ++-defn)
17
- open import Data.List.Effectful as Listₑ using (monad)
16
+ open import Data.List.Effectful as List using (monad)
18
17
open import Data.List.Relation.Unary.Any as Any using (Any; here; there)
19
18
open import Data.List.Membership.Propositional
20
19
open import Data.List.Membership.Propositional.Properties.Core
@@ -624,18 +623,16 @@ reverseAcc⁺ acc (x ∷ xs) (inj₂ (there y)) = reverseAcc⁺ (x ∷ acc) xs (
624
623
625
624
reverseAcc⁻ : ∀ acc xs → Any P (reverseAcc acc xs) → Any P acc ⊎ Any P xs
626
625
reverseAcc⁻ acc [] ps = inj₁ ps
627
- reverseAcc⁻ acc (x ∷ xs) ps rewrite ʳ++-defn xs {x ∷ acc} with ++⁻ (reverseAcc [] xs) ps
628
- ... | inj₂ (here p') = inj₂ (here p')
629
- ... | inj₂ (there ps') = inj₁ ps'
630
- ... | inj₁ ps' with reverseAcc⁻ [] xs ps'
631
- ... | inj₂ ps'' = inj₂ (there ps'')
626
+ reverseAcc⁻ acc (x ∷ xs) ps with reverseAcc⁻ (x ∷ acc) xs ps
627
+ ... | inj₁ (here px) = inj₂ (here px)
628
+ ... | inj₁ (there pxs) = inj₁ pxs
629
+ ... | inj₂ pxs = inj₂ (there pxs)
632
630
633
631
reverse⁺ : Any P xs → Any P (reverse xs)
634
632
reverse⁺ ps = reverseAcc⁺ [] _ (inj₂ ps)
635
633
636
634
reverse⁻ : Any P (reverse xs) → Any P xs
637
- reverse⁻ ps with reverseAcc⁻ [] _ ps
638
- ... | inj₂ ps' = ps'
635
+ reverse⁻ ps with inj₂ pxs ← reverseAcc⁻ [] _ ps = pxs
639
636
640
637
------------------------------------------------------------------------
641
638
-- pure
@@ -686,7 +683,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where
686
683
Any (λ f → Any (P ∘ f) xs) fs ↔⟨ Any-cong (λ _ → Any-cong (λ _ → pure↔) (_ ∎)) (_ ∎) ⟩
687
684
Any (λ f → Any (Any P ∘ pure ∘ f) xs) fs ↔⟨ Any-cong (λ _ → >>=↔ ) (_ ∎) ⟩
688
685
Any (λ f → Any P (xs >>= pure ∘ f)) fs ↔⟨ >>=↔ ⟩
689
- Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡⟨ cong (Any P) (Listₑ .Applicative.unfold-⊛ fs xs) ⟨
686
+ Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡⟨ cong (Any P) (List .Applicative.unfold-⊛ fs xs) ⟨
690
687
Any P (fs ⊛ xs) ∎
691
688
where open Related.EquationalReasoning
692
689
@@ -706,7 +703,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where
706
703
Any (λ x → Any (λ y → P (x , y)) ys) xs ↔⟨ pure↔ ⟩
707
704
Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (pure _,_) ↔⟨ ⊛↔ ⟩
708
705
Any (λ x, → Any (P ∘ x,) ys) (pure _,_ ⊛ xs) ↔⟨ ⊛↔ ⟩
709
- Any P (pure _,_ ⊛ xs ⊛ ys) ≡⟨ cong (Any P ∘′ (_⊛ ys)) (Listₑ .Applicative.unfold-<$> _,_ xs) ⟨
706
+ Any P (pure _,_ ⊛ xs ⊛ ys) ≡⟨ cong (Any P ∘′ (_⊛ ys)) (List .Applicative.unfold-<$> _,_ xs) ⟨
710
707
Any P (xs ⊗ ys) ∎
711
708
where open Related.EquationalReasoning
712
709
0 commit comments