Skip to content

Commit 157d8bb

Browse files
authored
Simplify Data.List.Relation.Unary.Any.Properties: remove dependency on List.Properties (#2323)
* simplifed proof; removed dependency on `List.Properties` * corrected module long name in comment
1 parent 098a65e commit 157d8bb

File tree

2 files changed

+10
-13
lines changed

2 files changed

+10
-13
lines changed

src/Data/List/Membership/Propositional/Properties/Core.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -5,8 +5,8 @@
55
------------------------------------------------------------------------
66

77
-- This file is needed to break the cyclic dependency with the proof
8-
-- `Any-cong` in `Data.Any.Properties` which relies on `Any↔` in this
9-
-- file.
8+
-- `Any-cong` in `Data.List.Relation.Unary.Any.Properties` which relies
9+
-- on `Any↔` defined in this file.
1010

1111
{-# OPTIONS --cubical-compatible --safe #-}
1212

src/Data/List/Relation/Unary/Any/Properties.agda

Lines changed: 8 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -13,8 +13,7 @@ open import Data.Bool.Properties using (T-∨; T-≡)
1313
open import Data.Empty using (⊥)
1414
open import Data.Fin.Base using (Fin; zero; suc)
1515
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)
1817
open import Data.List.Relation.Unary.Any as Any using (Any; here; there)
1918
open import Data.List.Membership.Propositional
2019
open import Data.List.Membership.Propositional.Properties.Core
@@ -624,18 +623,16 @@ reverseAcc⁺ acc (x ∷ xs) (inj₂ (there y)) = reverseAcc⁺ (x ∷ acc) xs (
624623

625624
reverseAcc⁻ : acc xs Any P (reverseAcc acc xs) Any P acc ⊎ Any P xs
626625
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)
632630

633631
reverse⁺ : Any P xs Any P (reverse xs)
634632
reverse⁺ ps = reverseAcc⁺ [] _ (inj₂ ps)
635633

636634
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
639636

640637
------------------------------------------------------------------------
641638
-- pure
@@ -686,7 +683,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where
686683
Any (λ f Any (P ∘ f) xs) fs ↔⟨ Any-cong (λ _ Any-cong (λ _ pure↔) (_ ∎)) (_ ∎) ⟩
687684
Any (λ f Any (Any P ∘ pure ∘ f) xs) fs ↔⟨ Any-cong (λ _ >>=↔ ) (_ ∎) ⟩
688685
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) ⟨
690687
Any P (fs ⊛ xs) ∎
691688
where open Related.EquationalReasoning
692689

@@ -706,7 +703,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where
706703
Any (λ x Any (λ y P (x , y)) ys) xs ↔⟨ pure↔ ⟩
707704
Any (λ _,_ Any (λ x Any (λ y P (x , y)) ys) xs) (pure _,_) ↔⟨ ⊛↔ ⟩
708705
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) ⟨
710707
Any P (xs ⊗ ys) ∎
711708
where open Related.EquationalReasoning
712709

0 commit comments

Comments
 (0)