@@ -245,16 +245,15 @@ commutativeMonoid {a} k A = record
245
245
246
246
empty-unique : ∀ {k} → xs ∼[ ⌊ k ⌋→ ] [] → xs ≡ []
247
247
empty-unique {xs = []} _ = refl
248
- empty-unique {xs = _ ∷ _} ∷∼[] with ⇒→ ∷∼[] (here refl)
249
- ... | ()
248
+ empty-unique {xs = _ ∷ _} ∷∼[] with () ← ⇒→ ∷∼[] (here refl)
250
249
251
250
-- _++_ is idempotent (under set equality).
252
251
253
252
++-idempotent : Idempotent {A = List A} _∼[ set ]_ _++_
254
253
++-idempotent xs {x} =
255
- x ∈ xs ++ xs ∼⟨ mk⇔ ([ id , id ]′ ∘ (Inverse.from $ ++↔))
256
- (Inverse.to ++↔ ∘ inj₁) ⟩
257
- x ∈ xs ∎
254
+ x ∈ xs ++ xs
255
+ ∼⟨ mk⇔ ([ id , id ]′ ∘ (Inverse.from $ ++↔)) (Inverse.to ++↔ ∘ inj₁) ⟩
256
+ x ∈ xs ∎
258
257
where open Related.EquationalReasoning
259
258
260
259
-- The list monad's bind distributes from the left over _++_.
@@ -538,62 +537,60 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys =
538
537
539
538
lemma : ∀ {xs ys} (inv : x ∷ xs ∼[ bag ] x ∷ ys) {z} →
540
539
Well-behaved (Inverse.to (∼→⊎↔⊎ inv {z}))
541
- lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ with
542
- zero ≡⟨⟩
543
- index-of {xs = x ∷ xs} (here p) ≡⟨⟩
544
- index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ ≡.sym $
545
- to-from (∼→⊎↔⊎ inv) {x = inj₁ p} hyp₂ ⟩
546
- index-of {xs = x ∷ xs} (to (∷↔ _) $ (from (∼→⊎↔⊎ inv) $ inj₁ q)) ≡⟨ ≡.cong index-of $
547
- strictlyInverseˡ (∷↔ _) (from inv (here q)) ⟩
548
- index-of {xs = x ∷ xs} (to (SK-sym inv) $ here q) ≡⟨ index-equality-preserved (SK-sym inv) refl ⟩
549
- index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p) ≡⟨ ≡.cong index-of $ ≡.sym $
550
- strictlyInverseˡ (∷↔ _) (from inv (here p)) ⟩
551
- index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p)) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $
552
- to-from (∼→⊎↔⊎ inv) {x = inj₂ z∈xs} hyp₁ ⟩
553
- index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₂ z∈xs) ≡⟨⟩
554
- index-of {xs = x ∷ xs} (there z∈xs) ≡⟨⟩
555
- suc (index-of {xs = xs} z∈xs) ∎
540
+ lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ = case contra of λ ()
556
541
where
557
542
open Inverse
558
543
open ≡-Reasoning
559
- ... | ()
544
+ contra : zero ≡ suc (index-of {xs = xs} z∈xs)
545
+ contra = begin
546
+ zero
547
+ ≡⟨⟩
548
+ index-of {xs = x ∷ xs} (here p)
549
+ ≡⟨⟩
550
+ index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p)
551
+ ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ to-from (∼→⊎↔⊎ inv) {x = inj₁ p} hyp₂ ⟨
552
+ index-of {xs = x ∷ xs} (to (∷↔ _) $ (from (∼→⊎↔⊎ inv) $ inj₁ q))
553
+ ≡⟨ ≡.cong index-of $ strictlyInverseˡ (∷↔ _) (from inv (here q)) ⟩
554
+ index-of {xs = x ∷ xs} (to (SK-sym inv) $ here q)
555
+ ≡⟨ index-equality-preserved (SK-sym inv) refl ⟩
556
+ index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p)
557
+ ≡⟨ ≡.cong index-of $ strictlyInverseˡ (∷↔ _) (from inv (here p)) ⟨
558
+ index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p))
559
+ ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ to-from (∼→⊎↔⊎ inv) {x = inj₂ z∈xs} hyp₁ ⟩
560
+ index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₂ z∈xs)
561
+ ≡⟨⟩
562
+ index-of {xs = x ∷ xs} (there z∈xs)
563
+ ≡⟨⟩
564
+ suc (index-of {xs = xs} z∈xs)
565
+ ∎
566
+
560
567
561
568
------------------------------------------------------------------------
562
569
-- Relationships to other relations
563
570
564
571
↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_
565
- ↭⇒∼bag {A = A} xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys)
572
+ ↭⇒∼bag xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys)
566
573
where
567
574
to : ∀ {xs ys} → xs ↭ ys → v ∈ xs → v ∈ ys
568
- to xs↭ys = Any -resp-↭ {A = A} xs↭ys
575
+ to xs↭ys = ∈ -resp-↭ xs↭ys
569
576
570
577
from : ∀ {xs ys} → xs ↭ ys → v ∈ ys → v ∈ xs
571
- from xs↭ys = Any -resp-↭ (↭-sym xs↭ys)
578
+ from xs↭ys = ∈ -resp-↭ (↭-sym xs↭ys)
572
579
573
580
from∘to : ∀ {xs ys} (p : xs ↭ ys) (q : v ∈ xs) → from p (to p q) ≡ q
574
- from∘to refl v∈xs = refl
575
- from∘to (prep _ _) (here refl) = refl
576
- from∘to (prep _ p) (there v∈xs) = ≡.cong there (from∘to p v∈xs)
577
- from∘to (swap x y p) (here refl) = refl
578
- from∘to (swap x y p) (there (here refl)) = refl
579
- from∘to (swap x y p) (there (there v∈xs)) = ≡.cong (there ∘ there) (from∘to p v∈xs)
580
- from∘to (trans p₁ p₂) v∈xs
581
- rewrite from∘to p₂ (Any-resp-↭ p₁ v∈xs)
582
- | from∘to p₁ v∈xs = refl
581
+ from∘to = ∈-resp-[σ⁻¹∘σ]
583
582
584
583
to∘from : ∀ {xs ys} (p : xs ↭ ys) (q : v ∈ ys) → to p (from p q) ≡ q
585
- to∘from p with from∘to (↭-sym p)
586
- ... | res rewrite ↭-sym-involutive p = res
584
+ to∘from p with res ← from∘to (↭-sym p) rewrite ↭-sym-involutive p = res
587
585
588
586
∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A}
589
- ∼bag⇒↭ {A = A} {[]} eq with empty-unique (↔-sym eq)
590
- ... | refl = refl
591
- ∼bag⇒↭ {A = A} {x ∷ xs} eq with ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl))
592
- ... | zs₁ , zs₂ , p rewrite p = begin
593
- x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩
594
- x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩
595
- x ∷ (zs₁ ++ zs₂) ↭⟨ shift x zs₁ zs₂ ⟨
596
- zs₁ ++ x ∷ zs₂ ∎
597
- where
598
- open CommutativeMonoid (commutativeMonoid bag A)
599
- open PermutationReasoning
587
+ ∼bag⇒↭ {A = A} {[]} eq with refl ← empty-unique (↔-sym eq) = refl
588
+ ∼bag⇒↭ {A = A} {x ∷ xs} eq
589
+ with zs₁ , zs₂ , p ← ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl)) rewrite p = begin
590
+ x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩
591
+ x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩
592
+ x ∷ (zs₁ ++ zs₂) ↭⟨ shift x zs₁ zs₂ ⟨
593
+ zs₁ ++ x ∷ zs₂ ∎
594
+ where
595
+ open CommutativeMonoid (commutativeMonoid bag A)
596
+ open PermutationReasoning
0 commit comments