/usr/share/agda-stdlib/Data/Star/Properties.agda is in agda-stdlib 0.6-2.
This file is owned by root:root, with mode 0o644.
The actual contents of the file can be viewed below.
1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70 71 72 73 74 75 76 77 78 79 80 81 82 83 84 85 86 87 88 89 90 91 92 93 94 95 | ------------------------------------------------------------------------
-- The Agda standard library
--
-- Some properties related to Data.Star
------------------------------------------------------------------------
module Data.Star.Properties where
open import Data.Star
open import Function
open import Relation.Binary
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl; sym; cong; cong₂)
import Relation.Binary.PreorderReasoning as PreR
◅◅-assoc : ∀ {i t} {I : Set i} {T : Rel I t} {i j k l}
(xs : Star T i j) (ys : Star T j k)
(zs : Star T k l) →
(xs ◅◅ ys) ◅◅ zs ≡ xs ◅◅ (ys ◅◅ zs)
◅◅-assoc ε ys zs = refl
◅◅-assoc (x ◅ xs) ys zs = cong (_◅_ x) (◅◅-assoc xs ys zs)
gmap-id : ∀ {i t} {I : Set i} {T : Rel I t} {i j} (xs : Star T i j) →
gmap id id xs ≡ xs
gmap-id ε = refl
gmap-id (x ◅ xs) = cong (_◅_ x) (gmap-id xs)
gmap-∘ : ∀ {i t} {I : Set i} {T : Rel I t}
{j u} {J : Set j} {U : Rel J u}
{k v} {K : Set k} {V : Rel K v}
(f : J → K) (g : U =[ f ]⇒ V)
(f′ : I → J) (g′ : T =[ f′ ]⇒ U)
{i j} (xs : Star T i j) →
(gmap {U = V} f g ∘ gmap f′ g′) xs ≡ gmap (f ∘ f′) (g ∘ g′) xs
gmap-∘ f g f′ g′ ε = refl
gmap-∘ f g f′ g′ (x ◅ xs) = cong (_◅_ (g (g′ x))) (gmap-∘ f g f′ g′ xs)
gmap-◅◅ : ∀ {i t j u}
{I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u}
(f : I → J) (g : T =[ f ]⇒ U)
{i j k} (xs : Star T i j) (ys : Star T j k) →
gmap {U = U} f g (xs ◅◅ ys) ≡ gmap f g xs ◅◅ gmap f g ys
gmap-◅◅ f g ε ys = refl
gmap-◅◅ f g (x ◅ xs) ys = cong (_◅_ (g x)) (gmap-◅◅ f g xs ys)
gmap-cong : ∀ {i t j u}
{I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u}
(f : I → J) (g : T =[ f ]⇒ U) (g′ : T =[ f ]⇒ U) →
(∀ {i j} (x : T i j) → g x ≡ g′ x) →
∀ {i j} (xs : Star T i j) →
gmap {U = U} f g xs ≡ gmap f g′ xs
gmap-cong f g g′ eq ε = refl
gmap-cong f g g′ eq (x ◅ xs) = cong₂ _◅_ (eq x) (gmap-cong f g g′ eq xs)
fold-◅◅ : ∀ {i p} {I : Set i}
(P : Rel I p) (_⊕_ : Transitive P) (∅ : Reflexive P) →
(∀ {i j} (x : P i j) → ∅ ⊕ x ≡ x) →
(∀ {i j k l} (x : P i j) (y : P j k) (z : P k l) →
(x ⊕ y) ⊕ z ≡ x ⊕ (y ⊕ z)) →
∀ {i j k} (xs : Star P i j) (ys : Star P j k) →
fold P _⊕_ ∅ (xs ◅◅ ys) ≡ fold P _⊕_ ∅ xs ⊕ fold P _⊕_ ∅ ys
fold-◅◅ P _⊕_ ∅ left-unit assoc ε ys = sym (left-unit _)
fold-◅◅ P _⊕_ ∅ left-unit assoc (x ◅ xs) ys = begin
x ⊕ fold P _⊕_ ∅ (xs ◅◅ ys) ≡⟨ cong (_⊕_ x) $
fold-◅◅ P _⊕_ ∅ left-unit assoc xs ys ⟩
x ⊕ (fold P _⊕_ ∅ xs ⊕ fold P _⊕_ ∅ ys) ≡⟨ sym (assoc x _ _) ⟩
(x ⊕ fold P _⊕_ ∅ xs) ⊕ fold P _⊕_ ∅ ys ∎
where open PropEq.≡-Reasoning
-- Reflexive transitive closures are preorders.
preorder : ∀ {i t} {I : Set i} (T : Rel I t) → Preorder _ _ _
preorder T = record
{ _≈_ = _≡_
; _∼_ = Star T
; isPreorder = record
{ isEquivalence = PropEq.isEquivalence
; reflexive = reflexive
; trans = _◅◅_
}
}
where
reflexive : _≡_ ⇒ Star T
reflexive refl = ε
-- Preorder reasoning for Star.
module StarReasoning {i t} {I : Set i} (T : Rel I t) where
open PreR (preorder T) public
renaming (_∼⟨_⟩_ to _⟶⋆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
infixr 2 _⟶⟨_⟩_
_⟶⟨_⟩_ : ∀ x {y z} → T x y → y IsRelatedTo z → x IsRelatedTo z
x ⟶⟨ x⟶y ⟩ y⟶⋆z = x ⟶⋆⟨ x⟶y ◅ ε ⟩ y⟶⋆z
|