Values.agda 43 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772
  1. module Issue784.Values where
  2. open import Data.Bool using (Bool; true; false; not)
  3. open import Data.String public using (String)
  4. open import Data.String.Properties public using (_≟_)
  5. open import Function public
  6. open import Data.List using (List; []; _∷_; _++_; [_]; boolFilter) renaming (map to mapL)
  7. open import Data.List.Relation.Unary.Any public using (Any; here; there) renaming (map to mapA; any to anyA)
  8. open import Data.Product public using (Σ; Σ-syntax; proj₁; proj₂; _,_; _×_) renaming (map to mapΣ)
  9. open import Data.Unit public using (⊤)
  10. open import Data.Unit.NonEta public using (Unit; unit)
  11. open import Data.Empty public using (⊥; ⊥-elim)
  12. open import Relation.Binary.Core hiding (_⇔_) public
  13. open import Relation.Nullary hiding (Irrelevant) public
  14. import Level
  15. open import Relation.Binary.PropositionalEquality public hiding ([_]) renaming (cong to ≡-cong; cong₂ to ≡-cong₂)
  16. open import Relation.Binary.PropositionalEquality.Core public renaming (sym to ≡-sym; trans to ≡-trans)
  17. {-
  18. ≢-sym : ∀ {ℓ} {A : Set ℓ} {x y : A} → x ≢ y → y ≢ x
  19. ≢-sym x≢y = x≢y ∘ ≡-sym
  20. -}
  21. ≡-elim : ∀ {ℓ} {X Y : Set ℓ} → X ≡ Y → X → Y
  22. ≡-elim refl p = p
  23. ≡-elim′ : ∀ {a ℓ} {A : Set a} {x y : A} → (P : A → Set ℓ) → x ≡ y → P x → P y
  24. ≡-elim′ P = ≡-elim ∘ (≡-cong P)
  25. Named : ∀ {ℓ} (A : Set ℓ) → Set ℓ
  26. Named A = String × A
  27. NamedType : ∀ ℓ → Set (Level.suc ℓ)
  28. NamedType ℓ = Named (Set ℓ)
  29. NamedValue : ∀ ℓ → Set (Level.suc ℓ)
  30. NamedValue ℓ = Named (Σ[ A ∈ Set ℓ ] A)
  31. Names : Set
  32. Names = List String
  33. Types : ∀ ℓ → Set (Level.suc ℓ)
  34. Types ℓ = List (NamedType ℓ)
  35. Values : ∀ ℓ → Set (Level.suc ℓ)
  36. Values ℓ = List (NamedValue ℓ)
  37. names : ∀ {ℓ} {A : Set ℓ} → List (Named A) → Names
  38. names = mapL proj₁
  39. types : ∀ {ℓ} → Values ℓ → Types ℓ
  40. types = mapL (mapΣ id proj₁)
  41. infix 4 _∈_ _∉_
  42. _∈_ : ∀ {ℓ} {A : Set ℓ} → A → List A → Set ℓ
  43. x ∈ xs = Any (_≡_ x) xs
  44. _∉_ : ∀ {ℓ} {A : Set ℓ} → A → List A → Set ℓ
  45. x ∉ xs = ¬ x ∈ xs
  46. x∉y∷l⇒x≢y : ∀ {ℓ} {A : Set ℓ} {x y : A} {l : List A} → x ∉ y ∷ l → x ≢ y
  47. x∉y∷l⇒x≢y x∉y∷l x≡y = x∉y∷l $ here x≡y
  48. x∉y∷l⇒x∉l : ∀ {ℓ} {A : Set ℓ} {x y : A} {l : List A} → x ∉ y ∷ l → x ∉ l
  49. x∉y∷l⇒x∉l x∉y∷l x∈l = x∉y∷l $ there x∈l
  50. x≢y⇒x∉l⇒x∉y∷l : ∀ {ℓ} {A : Set ℓ} {x y : A} {l : List A} → x ≢ y → x ∉ l → x ∉ y ∷ l
  51. x≢y⇒x∉l⇒x∉y∷l x≢y x∉l (here x≡y) = x≢y x≡y
  52. x≢y⇒x∉l⇒x∉y∷l x≢y x∉l (there x∈l) = x∉l x∈l
  53. infix 5 _∋!_ _∈?_
  54. _∈?_ : (s : String) (n : Names) → Dec (s ∈ n)
  55. s ∈? [] = no λ()
  56. s ∈? (h ∷ t) with s ≟ h
  57. ... | yes s≡h = yes $ here s≡h
  58. ... | no s≢h with s ∈? t
  59. ... | yes s∈t = yes $ there s∈t
  60. ... | no s∉t = no $ x≢y⇒x∉l⇒x∉y∷l s≢h s∉t
  61. _∋!_ : Names → String → Bool
  62. l ∋! e with e ∈? l
  63. ... | yes _ = true
  64. ... | no _ = false
  65. infix 4 _⊆_ _⊈_
  66. _⊆_ : ∀ {ℓ} {A : Set ℓ} → List A → List A → Set ℓ
  67. xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
  68. _⊈_ : ∀ {ℓ} {A : Set ℓ} → List A → List A → Set ℓ
  69. xs ⊈ ys = ¬ xs ⊆ ys
  70. infixl 5 _∪_
  71. _∪_ : ∀ {ℓ} {A : Set ℓ} → List A → List A → List A
  72. _∪_ = _++_
  73. infixl 6 _∩_
  74. _∩_ : Names → Names → Names
  75. x ∩ y = boolFilter (_∋!_ y) x
  76. infixl 6 _∖_ _∖∖_
  77. _∖_ : Names → Names → Names
  78. _∖_ x y = boolFilter (not ∘ _∋!_ y) x
  79. _∖∖_ : ∀ {ℓ} {A : Set ℓ} → List (Named A) → Names → List (Named A)
  80. _∖∖_ l n = boolFilter (not ∘ _∋!_ n ∘ proj₁) l
  81. filter-∈ : ∀ {ℓ} {A : Set ℓ} → List (Named A) → Names → List (Named A)
  82. filter-∈ l n = boolFilter (_∋!_ n ∘ proj₁) l
  83. infixr 5 _∷_
  84. data NonRepetitive {ℓ} {A : Set ℓ} : List A → Set ℓ where
  85. [] : NonRepetitive []
  86. _∷_ : ∀ {x xs} → x ∉ xs → NonRepetitive xs → NonRepetitive (x ∷ xs)
  87. infix 4 _≋_
  88. data _≋_ {ℓ} {A : Set ℓ} : List A → List A → Set ℓ where
  89. refl : ∀ {l} → l ≋ l
  90. perm : ∀ {l} l₁ x₁ x₂ l₂ → l ≋ l₁ ++ x₂ ∷ x₁ ∷ l₂ → l ≋ l₁ ++ x₁ ∷ x₂ ∷ l₂
  91. NonRepetitiveNames : ∀ {ℓ} {A : Set ℓ} → List (Named A) → Set
  92. NonRepetitiveNames = NonRepetitive ∘ names
  93. NonRepetitiveTypes : ∀ ℓ → Set (Level.suc ℓ)
  94. NonRepetitiveTypes ℓ = Σ[ t ∈ Types ℓ ] NonRepetitiveNames t
  95. -- lemmas
  96. ∪-assoc : ∀ {l} {A : Set l} (x y z : List A) → (x ∪ y) ∪ z ≡ x ∪ (y ∪ z)
  97. ∪-assoc [] y z = refl
  98. ∪-assoc (x ∷ xs) y z = ≡-elim′ (λ e → x ∷ e ≡ (x ∷ xs) ∪ (y ∪ z)) (≡-sym $ ∪-assoc xs y z) refl
  99. ≡⇒≋ : ∀ {ℓ} {A : Set ℓ} {x y : List A} → x ≡ y → x ≋ y
  100. ≡⇒≋ refl = refl
  101. private
  102. ∈-perm : ∀ {ℓ} {A : Set ℓ} {x : A} (l₁ : List A) (e₁ : A) (e₂ : A) (l₂ : List A) → x ∈ l₁ ++ e₁ ∷ e₂ ∷ l₂ → x ∈ l₁ ++ e₂ ∷ e₁ ∷ l₂
  103. ∈-perm [] e₁ e₂ l₂ (here .{e₁} .{e₂ ∷ l₂} px) = there $ here px
  104. ∈-perm [] e₁ e₂ l₂ (there .{e₁} .{e₂ ∷ l₂} (here .{e₂} .{l₂} px)) = here px
  105. ∈-perm [] e₁ e₂ l₂ (there .{e₁} .{e₂ ∷ l₂} (there .{e₂} .{l₂} pxs)) = there $ there pxs
  106. ∈-perm (h₁ ∷ t₁) e₁ e₂ l₂ (here .{h₁} .{t₁ ++ e₁ ∷ e₂ ∷ l₂} px) = here px
  107. ∈-perm (h₁ ∷ t₁) e₁ e₂ l₂ (there .{h₁} .{t₁ ++ e₁ ∷ e₂ ∷ l₂} pxs) = there $ ∈-perm t₁ e₁ e₂ l₂ pxs
  108. ≋⇒⊆ : ∀ {ℓ} {A : Set ℓ} {x y : List A} → x ≋ y → x ⊆ y
  109. ≋⇒⊆ refl p = p
  110. ≋⇒⊆ {x = x} {y = .(l₁ ++ x₁ ∷ x₂ ∷ l₂)} (perm l₁ x₁ x₂ l₂ p) {e} e∈x = ∈-perm l₁ x₂ x₁ l₂ $ ≋⇒⊆ p e∈x
  111. ≋-trans : ∀ {l} {A : Set l} {x y z : List A} → x ≋ y → y ≋ z → x ≋ z
  112. ≋-trans p refl = p
  113. ≋-trans p₁ (perm l₁ x₁ x₂ l₂ p₂) = perm l₁ x₁ x₂ l₂ $ ≋-trans p₁ p₂
  114. ≋-sym : ∀ {l} {A : Set l} {x y : List A} → x ≋ y → y ≋ x
  115. ≋-sym refl = refl
  116. ≋-sym (perm l₁ x₁ x₂ l₂ p) = ≋-trans (perm l₁ x₂ x₁ l₂ refl) (≋-sym p)
  117. ≋-del-ins-r : ∀ {l} {A : Set l} (l₁ : List A) (x : A) (l₂ l₃ : List A) → (l₁ ++ x ∷ l₂ ++ l₃) ≋ (l₁ ++ l₂ ++ x ∷ l₃)
  118. ≋-del-ins-r l₁ x [] l₃ = refl
  119. ≋-del-ins-r l₁ x (h ∷ t) l₃ = ≋-trans p₀ p₅
  120. where p₀ : (l₁ ++ x ∷ h ∷ t ++ l₃) ≋ (l₁ ++ h ∷ x ∷ t ++ l₃)
  121. p₀ = perm l₁ h x (t ++ l₃) refl
  122. p₁ : ((l₁ ++ [ h ]) ++ x ∷ t ++ l₃) ≋ ((l₁ ++ [ h ]) ++ t ++ x ∷ l₃)
  123. p₁ = ≋-del-ins-r (l₁ ++ [ h ]) x t l₃
  124. p₂ : (l₁ ++ [ h ]) ++ t ++ x ∷ l₃ ≡ l₁ ++ h ∷ t ++ x ∷ l₃
  125. p₂ = ∪-assoc l₁ [ h ] (t ++ x ∷ l₃)
  126. p₃ : (l₁ ++ [ h ]) ++ x ∷ t ++ l₃ ≡ l₁ ++ h ∷ x ∷ t ++ l₃
  127. p₃ = ∪-assoc l₁ [ h ] (x ∷ t ++ l₃)
  128. p₄ : ((l₁ ++ [ h ]) ++ x ∷ t ++ l₃) ≋ (l₁ ++ h ∷ t ++ x ∷ l₃)
  129. p₄ = ≡-elim′ (λ y → ((l₁ ++ [ h ]) ++ x ∷ t ++ l₃) ≋ y) p₂ p₁
  130. p₅ : (l₁ ++ h ∷ x ∷ t ++ l₃) ≋ (l₁ ++ h ∷ t ++ x ∷ l₃)
  131. p₅ = ≡-elim′ (λ y → y ≋ (l₁ ++ h ∷ t ++ x ∷ l₃)) p₃ p₄
  132. ≋-del-ins-l : ∀ {l} {A : Set l} (l₁ l₂ : List A) (x : A) (l₃ : List A) → (l₁ ++ l₂ ++ x ∷ l₃) ≋ (l₁ ++ x ∷ l₂ ++ l₃)
  133. ≋-del-ins-l l₁ l₂ x l₃ = ≋-sym $ ≋-del-ins-r l₁ x l₂ l₃
  134. x∪[]≡x : ∀ {ℓ} {A : Set ℓ} (x : List A) → x ∪ [] ≡ x
  135. x∪[]≡x [] = refl
  136. x∪[]≡x (h ∷ t) = ≡-trans p₀ p₁
  137. where p₀ : (h ∷ t) ++ [] ≡ h ∷ t ++ []
  138. p₀ = ∪-assoc [ h ] t []
  139. p₁ : h ∷ t ++ [] ≡ h ∷ t
  140. p₁ = ≡-cong (λ x → h ∷ x) (x∪[]≡x t)
  141. x∖[]≡x : (x : Names) → x ∖ [] ≡ x
  142. x∖[]≡x [] = refl
  143. x∖[]≡x (h ∷ t) = ≡-cong (_∷_ h) (x∖[]≡x t)
  144. t∖[]≡t : ∀ {ℓ} {A : Set ℓ} (t : List (Named A)) → t ∖∖ [] ≡ t
  145. t∖[]≡t [] = refl
  146. t∖[]≡t (h ∷ t) = ≡-cong (_∷_ h) (t∖[]≡t t)
  147. e∷x≋e∷y : ∀ {ℓ} {A : Set ℓ} (e : A) {x y : List A} → x ≋ y → e ∷ x ≋ e ∷ y
  148. e∷x≋e∷y _ refl = refl
  149. e∷x≋e∷y x (perm l₁ e₁ e₂ l₂ p) = perm (x ∷ l₁) e₁ e₂ l₂ (e∷x≋e∷y x p)
  150. ∪-sym : ∀ {ℓ} {A : Set ℓ} (x y : List A) → x ∪ y ≋ y ∪ x
  151. ∪-sym [] y = ≡-elim′ (λ z → y ≋ z) (≡-sym $ x∪[]≡x y) refl
  152. ∪-sym x [] = ≡-elim′ (λ z → z ≋ x) (≡-sym $ x∪[]≡x x) refl
  153. ∪-sym x (y ∷ ys) = ≡-elim′ (λ z → x ++ (y ∷ ys) ≋ z) p₃ (≋-trans p₁ p₂) where
  154. p₁ : x ++ (y ∷ ys) ≋ y ∷ (x ++ ys)
  155. p₁ = ≋-del-ins-l [] x y ys
  156. p₂ : y ∷ (x ++ ys) ≋ y ∷ (ys ++ x)
  157. p₂ = e∷x≋e∷y y $ ∪-sym x ys
  158. p₃ : y ∷ (ys ++ x) ≡ (y ∷ ys) ++ x
  159. p₃ = ∪-assoc [ y ] ys x
  160. y≋ỳ⇒x∪y≋x∪ỳ : ∀ {ℓ} {A : Set ℓ} (x : List A) {y ỳ : List A} → y ≋ ỳ → x ∪ y ≋ x ∪ ỳ
  161. y≋ỳ⇒x∪y≋x∪ỳ [] p = p
  162. y≋ỳ⇒x∪y≋x∪ỳ (h ∷ t) p = e∷x≋e∷y h (y≋ỳ⇒x∪y≋x∪ỳ t p)
  163. x≋x̀⇒x∪y≋x̀∪y : ∀ {ℓ} {A : Set ℓ} {x x̀ : List A} → x ≋ x̀ → (y : List A) → x ∪ y ≋ x̀ ∪ y
  164. x≋x̀⇒x∪y≋x̀∪y {x = x} {x̀ = x̀} p y = ≋-trans (∪-sym x y) $ ≋-trans (y≋ỳ⇒x∪y≋x∪ỳ y p) (∪-sym y x̀)
  165. x⊆y≋z : ∀ {ℓ} {A : Set ℓ} {x y z : List A} → x ⊆ y → y ≋ z → x ⊆ z
  166. x⊆y≋z x⊆y refl = x⊆y
  167. x⊆y≋z {x = x} {y = y} {z = .(l₁ ++ x₁ ∷ x₂ ∷ l₂)} x⊆y (perm l₁ x₁ x₂ l₂ p) = ∈-perm l₁ x₂ x₁ l₂ ∘ x⊆y≋z x⊆y p
  168. x≋y⊆z : ∀ {ℓ} {A : Set ℓ} {x y z : List A} → x ≋ y → x ⊆ z → y ⊆ z
  169. x≋y⊆z refl x⊆z = x⊆z
  170. x≋y⊆z {y = .(l₁ ++ x₁ ∷ x₂ ∷ l₂)} {z = z} (perm l₁ x₁ x₂ l₂ p) x⊆z = x≋y⊆z p x⊆z ∘ ∈-perm l₁ x₁ x₂ l₂
  171. x⊆x∪y : ∀ {ℓ} {A : Set ℓ} (x y : List A) → x ⊆ x ∪ y
  172. x⊆x∪y .(x ∷ xs) y (here {x} {xs} px) = here px
  173. x⊆x∪y .(x ∷ xs) y (there {x} {xs} pxs) = there $ x⊆x∪y xs y pxs
  174. x∪y⊆z⇒x⊆z : ∀ {ℓ} {A : Set ℓ} (x y : List A) {z : List A} → x ∪ y ⊆ z → x ⊆ z
  175. x∪y⊆z⇒x⊆z x y x∪y⊆z = x∪y⊆z ∘ x⊆x∪y x y
  176. x∪y⊆z⇒y⊆z : ∀ {ℓ} {A : Set ℓ} (x y : List A) {z : List A} → x ∪ y ⊆ z → y ⊆ z
  177. x∪y⊆z⇒y⊆z x y = x∪y⊆z⇒x⊆z y x ∘ x≋y⊆z (∪-sym x y)
  178. n-x∪y : ∀ {ℓ} {A : Set ℓ} (x y : List (Named A)) → names (x ∪ y) ≡ names x ∪ names y
  179. n-x∪y [] _ = refl
  180. n-x∪y {ℓ} (x ∷ xs) y = ≡-trans p₁ p₂ where
  181. nx = [ proj₁ x ]
  182. nxs = names xs
  183. ny = names y
  184. p₁ : nx ∪ names (xs ∪ y) ≡ nx ∪ (nxs ∪ ny)
  185. p₁ = ≡-cong (λ z → nx ∪ z) (n-x∪y xs y)
  186. p₂ : nx ∪ (nxs ∪ ny) ≡ (nx ∪ nxs) ∪ ny
  187. p₂ = ≡-sym $ ∪-assoc nx nxs ny
  188. t-x∪y : ∀ {ℓ} (x y : Values ℓ) → types (x ∪ y) ≡ types x ∪ types y
  189. t-x∪y [] _ = refl
  190. t-x∪y (x ∷ xs) y = ≡-trans p₁ p₂ where
  191. nx = types [ x ]
  192. nxs = types xs
  193. ny = types y
  194. p₁ : nx ∪ types (xs ∪ y) ≡ nx ∪ (nxs ∪ ny)
  195. p₁ = ≡-cong (λ z → nx ∪ z) (t-x∪y xs y)
  196. p₂ : nx ∪ (nxs ∪ ny) ≡ (nx ∪ nxs) ∪ ny
  197. p₂ = ≡-sym $ ∪-assoc nx nxs ny
  198. n-x≋y : ∀ {ℓ} {A : Set ℓ} {x y : List (Named A)} → x ≋ y → names x ≋ names y
  199. n-x≋y refl = refl
  200. n-x≋y (perm {x} l₁ e₁ e₂ l₂ p) = p₃ where
  201. n-l₁e₁e₂l₂ : ∀ e₁ e₂ → names (l₁ ++ e₁ ∷ e₂ ∷ l₂) ≡ names l₁ ++ names [ e₁ ] ++ names [ e₂ ] ++ names l₂
  202. n-l₁e₁e₂l₂ e₁ e₂ = ≡-trans p₁ $ ≡-trans p₂ p₃ where
  203. p₁ : names (l₁ ++ e₁ ∷ e₂ ∷ l₂) ≡ names l₁ ++ names (e₁ ∷ e₂ ∷ l₂)
  204. p₁ = n-x∪y l₁ (e₁ ∷ e₂ ∷ l₂)
  205. p₂ : names l₁ ++ names (e₁ ∷ e₂ ∷ l₂) ≡ names l₁ ++ names [ e₁ ] ++ names (e₂ ∷ l₂)
  206. p₂ = ≡-cong (λ z → names l₁ ++ z) (n-x∪y [ e₁ ] (e₂ ∷ l₂))
  207. p₃ : names l₁ ++ names [ e₁ ] ++ names (e₂ ∷ l₂) ≡ names l₁ ++ names [ e₁ ] ++ names [ e₂ ] ++ names l₂
  208. p₃ = ≡-cong (λ z → names l₁ ++ names [ e₁ ] ++ z) (n-x∪y [ e₂ ] l₂)
  209. p₁ : names x ≋ names l₁ ++ proj₁ e₂ ∷ proj₁ e₁ ∷ names l₂
  210. p₁ = ≡-elim′ (λ z → names x ≋ z) (n-l₁e₁e₂l₂ e₂ e₁) (n-x≋y p)
  211. p₂ : names x ≋ names l₁ ++ proj₁ e₁ ∷ proj₁ e₂ ∷ names l₂
  212. p₂ = perm (names l₁) (proj₁ e₁) (proj₁ e₂) (names l₂) p₁
  213. p₃ : names x ≋ names (l₁ ++ e₁ ∷ e₂ ∷ l₂)
  214. p₃ = ≡-elim′ (λ z → names x ≋ z) (≡-sym $ n-l₁e₁e₂l₂ e₁ e₂) p₂
  215. n-types : ∀ {ℓ} (x : Values ℓ) → names (types x) ≡ names x
  216. n-types [] = refl
  217. n-types (x ∷ xs) = ≡-cong (λ z → proj₁ x ∷ z) (n-types xs)
  218. nr-x≋y : ∀ {ℓ} {A : Set ℓ} {x y : List A} → x ≋ y → NonRepetitive x → NonRepetitive y
  219. nr-x≋y refl u = u
  220. nr-x≋y {y = .(l₁ ++ e₁ ∷ e₂ ∷ l₂)} (perm l₁ e₁ e₂ l₂ p) u = ≋-step l₁ e₂ e₁ l₂ (nr-x≋y p u) where
  221. ∉-step : ∀ {ℓ} {A : Set ℓ} {x : A} (l₁ : List A) (e₁ : A) (e₂ : A) (l₂ : List A) → x ∉ l₁ ++ e₁ ∷ e₂ ∷ l₂ → x ∉ l₁ ++ e₂ ∷ e₁ ∷ l₂
  222. ∉-step l₁ e₁ e₂ l₂ x∉l x∈l = ⊥-elim $ x∉l $ ∈-perm l₁ e₂ e₁ l₂ x∈l
  223. ≋-step : ∀ {ℓ} {A : Set ℓ} (l₁ : List A) (e₁ : A) (e₂ : A) (l₂ : List A) → NonRepetitive (l₁ ++ e₁ ∷ e₂ ∷ l₂) → NonRepetitive (l₁ ++ e₂ ∷ e₁ ∷ l₂)
  224. ≋-step [] e₁ e₂ l₂ (_∷_ .{e₁} .{e₂ ∷ l₂} e₁∉e₂∷l₂ (_∷_ .{e₂} .{l₂} e₂∉l₂ pU)) = e₂∉e₁∷l₂ ∷ e₁∉l₂ ∷ pU where
  225. e₁∉l₂ = e₁ ∉ l₂ ∋ x∉y∷l⇒x∉l e₁∉e₂∷l₂
  226. e₂∉e₁∷l₂ = e₂ ∉ e₁ ∷ l₂ ∋ x≢y⇒x∉l⇒x∉y∷l (≢-sym $ x∉y∷l⇒x≢y e₁∉e₂∷l₂) e₂∉l₂
  227. ≋-step (h₁ ∷ t₁) e₁ e₂ l₂ (_∷_ .{h₁} .{t₁ ++ e₁ ∷ e₂ ∷ l₂} p∉ pU) = ∉-step t₁ e₁ e₂ l₂ p∉ ∷ ≋-step t₁ e₁ e₂ l₂ pU
  228. nr-x⇒nr-t-x : ∀ {ℓ} {x : Values ℓ} → NonRepetitiveNames x → NonRepetitiveNames (types x)
  229. nr-x⇒nr-t-x {x = x} = ≡-elim′ NonRepetitive (≡-sym $ n-types x)
  230. n-x∖y : ∀ {ℓ} {A : Set ℓ} (x : List (Named A)) (y : Names) → names (x ∖∖ y) ≡ names x ∖ y
  231. n-x∖y [] _ = refl
  232. n-x∖y (x ∷ xs) ny with not $ ny ∋! proj₁ x
  233. ... | false = n-x∖y xs ny
  234. ... | true = ≡-trans p₁ p₂ where
  235. p₁ : names (x ∷ (xs ∖∖ ny)) ≡ proj₁ x ∷ names (xs ∖∖ ny)
  236. p₁ = n-x∪y [ x ] (xs ∖∖ ny)
  237. p₂ : proj₁ x ∷ names (xs ∖∖ ny) ≡ proj₁ x ∷ (names xs ∖ ny)
  238. p₂ = ≡-cong (λ z → proj₁ x ∷ z) (n-x∖y xs ny)
  239. t-x∖y : ∀ {ℓ} (x : Values ℓ) (y : Names) → types (x ∖∖ y) ≡ types x ∖∖ y
  240. t-x∖y [] _ = refl
  241. t-x∖y (x ∷ xs) ny with not $ ny ∋! proj₁ x
  242. ... | false = t-x∖y xs ny
  243. ... | true = ≡-trans p₁ p₂ where
  244. x̀ = types [ x ]
  245. p₁ : types (x ∷ (xs ∖∖ ny)) ≡ x̀ ∪ types (xs ∖∖ ny)
  246. p₁ = t-x∪y [ x ] (xs ∖∖ ny)
  247. p₂ : x̀ ∪ types (xs ∖∖ ny) ≡ x̀ ∪ (types xs ∖∖ ny)
  248. p₂ = ≡-cong (λ z → x̀ ∪ z) (t-x∖y xs ny)
  249. n-filter-∈ : ∀ {ℓ} {A : Set ℓ} (x : List (Named A)) (y : Names) → names (filter-∈ x y) ≡ names x ∩ y
  250. n-filter-∈ [] _ = refl
  251. n-filter-∈ (x ∷ xs) ny with ny ∋! proj₁ x
  252. ... | false = n-filter-∈ xs ny
  253. ... | true = ≡-trans p₁ p₂ where
  254. p₁ : names (x ∷ (filter-∈ xs ny)) ≡ proj₁ x ∷ names (filter-∈ xs ny)
  255. p₁ = n-x∪y [ x ] (filter-∈ xs ny)
  256. p₂ : proj₁ x ∷ names (filter-∈ xs ny) ≡ proj₁ x ∷ (names xs ∩ ny)
  257. p₂ = ≡-cong (λ z → proj₁ x ∷ z) (n-filter-∈ xs ny)
  258. t-filter-∈ : ∀ {ℓ} (x : Values ℓ) (y : Names) → types (filter-∈ x y) ≡ filter-∈ (types x) y
  259. t-filter-∈ [] _ = refl
  260. t-filter-∈ (x ∷ xs) ny with ny ∋! proj₁ x
  261. ... | false = t-filter-∈ xs ny
  262. ... | true = ≡-trans p₁ p₂ where
  263. x̀ = types [ x ]
  264. p₁ : types (x ∷ filter-∈ xs ny) ≡ x̀ ∪ types (filter-∈ xs ny)
  265. p₁ = t-x∪y [ x ] (filter-∈ xs ny)
  266. p₂ : x̀ ∪ types (filter-∈ xs ny) ≡ x̀ ∪ filter-∈ (types xs) ny
  267. p₂ = ≡-cong (λ z → x̀ ∪ z) (t-filter-∈ xs ny)
  268. []⊆x : ∀ {ℓ} {A : Set ℓ} (x : List A) → [] ⊆ x
  269. []⊆x _ ()
  270. ∀x∉[] : ∀ {ℓ} {A : Set ℓ} {x : A} → x ∉ []
  271. ∀x∉[] ()
  272. x⊆[]⇒x≡[] : ∀ {ℓ} {A : Set ℓ} {x : List A} → x ⊆ [] → x ≡ []
  273. x⊆[]⇒x≡[] {x = []} _ = refl
  274. x⊆[]⇒x≡[] {x = _ ∷ _} x⊆[] = ⊥-elim $ ∀x∉[] $ x⊆[] (here refl)
  275. x∩y⊆x : ∀ {ℓ} {A : Set ℓ} (x : List (Named A)) (y : Names) → filter-∈ x y ⊆ x
  276. x∩y⊆x [] _ = λ()
  277. x∩y⊆x (h ∷ t) y with y ∋! proj₁ h
  278. ... | false = there ∘ x∩y⊆x t y
  279. ... | true = f where
  280. f : h ∷ filter-∈ t y ⊆ h ∷ t
  281. f (here {x = .h} p) = here p
  282. f (there {xs = .(filter-∈ t y)} p) = there $ x∩y⊆x t y p
  283. e∈x⇒e∈y∪x : ∀ {ℓ} {A : Set ℓ} {e : A} {x : List A} (y : List A) → e ∈ x → e ∈ y ∪ x
  284. e∈x⇒e∈y∪x [] = id
  285. e∈x⇒e∈y∪x (h ∷ t) = there ∘ e∈x⇒e∈y∪x t
  286. e∈x⇒e∈x∪y : ∀ {ℓ} {A : Set ℓ} {e : A} {x : List A} (y : List A) → e ∈ x → e ∈ x ∪ y
  287. e∈x⇒e∈x∪y {e = e} {x = x} y e∈x = x⊆y≋z f (∪-sym y x) (e ∈ [ e ] ∋ here refl) where
  288. f : [ e ] ⊆ y ∪ x
  289. f {è} (here {x = .e} p) = ≡-elim′ (λ z → z ∈ y ∪ x) (≡-sym p) (e∈x⇒e∈y∪x y e∈x)
  290. f (there ())
  291. x∪y⊆x̀∪ỳ : ∀ {ℓ} {A : Set ℓ} {x x̀ y ỳ : List A} → x ⊆ x̀ → y ⊆ ỳ → x ∪ y ⊆ x̀ ∪ ỳ
  292. x∪y⊆x̀∪ỳ {x = []} {x̀ = []} _ y⊆ỳ = y⊆ỳ
  293. x∪y⊆x̀∪ỳ {x = []} {x̀ = _ ∷ t̀} _ y⊆ỳ = there ∘ x∪y⊆x̀∪ỳ ([]⊆x t̀) y⊆ỳ
  294. x∪y⊆x̀∪ỳ {x = h ∷ t} {x̀ = x̀} {y = y} {ỳ = ỳ} x⊆x̀ y⊆ỳ = f where
  295. f : (h ∷ t) ∪ y ⊆ x̀ ∪ ỳ
  296. f {e} (here {x = .h} e≡h) = e∈x⇒e∈x∪y ỳ (x⊆x̀ $ here e≡h)
  297. f {e} (there {xs = .(t ∪ y)} p) = x∪y⊆x̀∪ỳ (x∪y⊆z⇒y⊆z [ h ] t x⊆x̀) y⊆ỳ p
  298. x∖y⊆x : (x y : Names) → x ∖ y ⊆ x
  299. x∖y⊆x [] _ = λ()
  300. x∖y⊆x (h ∷ t) y with y ∋! h
  301. ... | true = there ∘ x∖y⊆x t y
  302. ... | false = x∪y⊆x̀∪ỳ (≋⇒⊆ $ ≡⇒≋ $ refl {x = [ h ]}) (x∖y⊆x t y)
  303. t≋t∖n∪t∩n : ∀ {ℓ} {A : Set ℓ} (t : List (Named A)) (n : Names) → t ≋ (t ∖∖ n) ∪ filter-∈ t n
  304. t≋t∖n∪t∩n [] _ = refl
  305. t≋t∖n∪t∩n (h ∷ t) n with n ∋! proj₁ h
  306. ... | false = e∷x≋e∷y h $ t≋t∖n∪t∩n t n
  307. ... | true = ≋-trans p₁ p₂ where
  308. p₁ : h ∷ t ≋ (h ∷ (t ∖∖ n)) ∪ filter-∈ t n
  309. p₁ = e∷x≋e∷y h $ t≋t∖n∪t∩n t n
  310. p₂ : (h ∷ (t ∖∖ n)) ∪ filter-∈ t n ≋ (t ∖∖ n) ∪ (h ∷ filter-∈ t n)
  311. p₂ = ≋-del-ins-r [] h (t ∖∖ n) (filter-∈ t n)
  312. e₁∈x⇒e₂∉x⇒e≢e₂ : ∀ {ℓ} {A : Set ℓ} {e₁ e₂ : A} {x : List A} → e₁ ∈ x → e₂ ∉ x → e₁ ≢ e₂
  313. e₁∈x⇒e₂∉x⇒e≢e₂ e₁∈x e₂∉x refl = e₂∉x e₁∈x
  314. e₁∈e₂∷x⇒e₁≢e₂⇒e₁∈x : ∀ {ℓ} {A : Set ℓ} {e₁ e₂ : A} {x : List A} → e₁ ∈ e₂ ∷ x → e₁ ≢ e₂ → e₁ ∈ x
  315. e₁∈e₂∷x⇒e₁≢e₂⇒e₁∈x (here e₁≡e₂) e₁≢e₂ = ⊥-elim $ e₁≢e₂ e₁≡e₂
  316. e₁∈e₂∷x⇒e₁≢e₂⇒e₁∈x (there e₁∈x) _ = e₁∈x
  317. x⊆e∷y⇒e∉x⇒x⊆y : ∀ {ℓ} {A : Set ℓ} {e : A} {x y : List A} → x ⊆ e ∷ y → e ∉ x → x ⊆ y
  318. x⊆e∷y⇒e∉x⇒x⊆y {e = e} {x = x} {y = y} x⊆e∷y e∉x {è} è∈x = e₁∈e₂∷x⇒e₁≢e₂⇒e₁∈x (x⊆e∷y è∈x) (e₁∈x⇒e₂∉x⇒e≢e₂ è∈x e∉x)
  319. n-e∉l⇒e∉l : ∀ {ℓ} {A : Set ℓ} {e : Named A} {l : List (Named A)} → proj₁ e ∉ names l → e ∉ l
  320. n-e∉l⇒e∉l {e = e} n-e∉l (here {x = è} p) = ⊥-elim $ n-e∉l $ here $ ≡-cong proj₁ p
  321. n-e∉l⇒e∉l n-e∉l (there p) = n-e∉l⇒e∉l (x∉y∷l⇒x∉l n-e∉l) p
  322. nr-names⇒nr : ∀ {ℓ} {A : Set ℓ} {l : List (Named A)} → NonRepetitiveNames l → NonRepetitive l
  323. nr-names⇒nr {l = []} [] = []
  324. nr-names⇒nr {l = _ ∷ _} (nh∉nt ∷ nr-t) = n-e∉l⇒e∉l nh∉nt ∷ nr-names⇒nr nr-t
  325. e∉x⇒e∉y⇒e∉x∪y : ∀ {ℓ} {A : Set ℓ} {e : A} {x y : List A} → e ∉ x → e ∉ y → e ∉ x ∪ y
  326. e∉x⇒e∉y⇒e∉x∪y {x = []} _ e∉y e∈y = ⊥-elim $ e∉y e∈y
  327. e∉x⇒e∉y⇒e∉x∪y {x = h ∷ t} e∉x e∉y (here e≡h) = ⊥-elim $ e∉x $ here e≡h
  328. e∉x⇒e∉y⇒e∉x∪y {x = h ∷ t} e∉x e∉y (there e∈t∪y) = e∉x⇒e∉y⇒e∉x∪y (x∉y∷l⇒x∉l e∉x) e∉y e∈t∪y
  329. nr-x∖y∪y : {x y : Names} → NonRepetitive x → NonRepetitive y → NonRepetitive (x ∖ y ∪ y)
  330. nr-x∖y∪y {x = []} _ nr-y = nr-y
  331. nr-x∖y∪y {x = x ∷ xs} {y = y} (x∉xs ∷ nr-xs) nr-y with x ∈? y
  332. ... | yes x∈y = nr-x∖y∪y nr-xs nr-y
  333. ... | no x∉y = e∉x⇒e∉y⇒e∉x∪y (⊥-elim ∘ x∉xs ∘ x∖y⊆x xs y) x∉y ∷ nr-x∖y∪y nr-xs nr-y
  334. e∉l∖[e] : (e : String) (l : Names) → e ∉ l ∖ [ e ]
  335. e∉l∖[e] _ [] = λ()
  336. e∉l∖[e] e (h ∷ t) with h ≟ e
  337. ... | yes _ = e∉l∖[e] e t
  338. ... | no h≢e = x≢y⇒x∉l⇒x∉y∷l (≢-sym h≢e) (e∉l∖[e] e t)
  339. e∉l⇒l∖e≡l : {e : String} {l : Names} → e ∉ l → l ∖ [ e ] ≡ l
  340. e∉l⇒l∖e≡l {e = e} {l = []} _ = refl
  341. e∉l⇒l∖e≡l {e = e} {l = h ∷ t} e∉l with h ∈? [ e ]
  342. ... | yes h∈e = ⊥-elim $ x≢y⇒x∉l⇒x∉y∷l (≢-sym $ x∉y∷l⇒x≢y e∉l) (λ()) h∈e
  343. ... | no h∉e = ≡-cong (_∷_ h) (e∉l⇒l∖e≡l $ x∉y∷l⇒x∉l e∉l)
  344. e∈l⇒nr-l⇒l∖e∪e≋l : {e : String} {l : Names} → e ∈ l → NonRepetitive l → l ∖ [ e ] ∪ [ e ] ≋ l
  345. e∈l⇒nr-l⇒l∖e∪e≋l {l = []} () _
  346. e∈l⇒nr-l⇒l∖e∪e≋l {e = e} {l = h ∷ t} e∈h∷t (h∉t ∷ nr-t) with h ∈? [ e ]
  347. ... | yes (here h≡e) = ≋-trans (≡⇒≋ $ ≡-trans p₁ p₂) (∪-sym t [ h ]) where
  348. p₁ : t ∖ [ e ] ∪ [ e ] ≡ t ∖ [ h ] ∪ [ h ]
  349. p₁ = ≡-cong (λ x → t ∖ [ x ] ∪ [ x ]) (≡-sym h≡e)
  350. p₂ : t ∖ [ h ] ∪ [ h ] ≡ t ∪ [ h ]
  351. p₂ = ≡-cong (λ x → x ∪ [ h ]) (e∉l⇒l∖e≡l h∉t)
  352. ... | yes (there ())
  353. ... | no h∉e with e∈h∷t
  354. ... | here e≡h = ⊥-elim ∘ h∉e ∘ here ∘ ≡-sym $ e≡h
  355. ... | there e∈t = e∷x≋e∷y _ $ e∈l⇒nr-l⇒l∖e∪e≋l e∈t nr-t
  356. nr-x∖y : {x : Names} → NonRepetitive x → (y : Names) → NonRepetitive (x ∖ y)
  357. nr-x∖y {x = []} _ _ = []
  358. nr-x∖y {x = x ∷ xs} (x∉xs ∷ nr-xs) y with x ∈? y
  359. ... | yes x∈y = nr-x∖y nr-xs y
  360. ... | no x∉y = ⊥-elim ∘ x∉xs ∘ x∖y⊆x xs y ∷ nr-x∖y nr-xs y
  361. x⊆y⇒e∉y⇒e∉x : ∀ {ℓ} {A : Set ℓ} {e : A} {x y : List A} → x ⊆ y → e ∉ y → e ∉ x
  362. x⊆y⇒e∉y⇒e∉x x⊆y e∉y e∈x = e∉y $ x⊆y e∈x
  363. e∈n-l⇒∃è,n-è≡e×è∈l : ∀ {ℓ} {A : Set ℓ} {e : String} {l : List (Named A)} → e ∈ names l → Σ[ è ∈ Named A ] (e ≡ proj₁ è × è ∈ l)
  364. e∈n-l⇒∃è,n-è≡e×è∈l {l = []} ()
  365. e∈n-l⇒∃è,n-è≡e×è∈l {l = h ∷ t} (here e≡n-h) = h , e≡n-h , here refl
  366. e∈n-l⇒∃è,n-è≡e×è∈l {l = h ∷ t} (there e∈n-t) with e∈n-l⇒∃è,n-è≡e×è∈l e∈n-t
  367. ... | è , n-è≡e , è∈l = è , n-è≡e , there è∈l
  368. x⊆y⇒nx⊆ny : ∀ {ℓ} {A : Set ℓ} {x y : List (Named A)} → x ⊆ y → names x ⊆ names y
  369. x⊆y⇒nx⊆ny {x = x} {y = y} x⊆y {e} e∈n-x with e ∈? names y
  370. ... | yes e∈n-y = e∈n-y
  371. ... | no e∉n-y with e∈n-l⇒∃è,n-è≡e×è∈l e∈n-x
  372. ... | è , e≡n-è , è∈x = ⊥-elim $ x⊆y⇒e∉y⇒e∉x x⊆y (n-e∉l⇒e∉l $ ≡-elim′ (λ z → z ∉ names y) e≡n-è e∉n-y) è∈x
  373. nr-x⇒nr-x∩y : ∀ {ℓ} {A : Set ℓ} {x : List (Named A)} → NonRepetitiveNames x → (y : Names) → NonRepetitiveNames (filter-∈ x y)
  374. nr-x⇒nr-x∩y {x = []} _ _ = []
  375. nr-x⇒nr-x∩y {x = h ∷ t} (h∉t ∷ nr-t) y with y ∋! proj₁ h
  376. ... | false = nr-x⇒nr-x∩y nr-t y
  377. ... | true = x⊆y⇒e∉y⇒e∉x (x⊆y⇒nx⊆ny $ x∩y⊆x t y) h∉t ∷ nr-x⇒nr-x∩y nr-t y
  378. e∈l⇒[e]⊆l : ∀ {ℓ} {A : Set ℓ} {e : A} {l : List A} → e ∈ l → [ e ] ⊆ l
  379. e∈l⇒[e]⊆l e∈l (here è≡e) = ≡-elim′ (λ x → x ∈ _) (≡-sym è≡e) e∈l
  380. e∈l⇒[e]⊆l e∈l (there ())
  381. a∈x⇒x≋y⇒a∈y : ∀ {ℓ} {A : Set ℓ} {x y : List A} {a : A} → a ∈ x → x ≋ y → a ∈ y
  382. a∈x⇒x≋y⇒a∈y a∈x x≋y = x⊆y≋z (e∈l⇒[e]⊆l a∈x) x≋y (here refl)
  383. x⊆z⇒y⊆z⇒x∪y⊆z : ∀ {ℓ} {A : Set ℓ} {x y z : List A} → x ⊆ z → y ⊆ z → x ∪ y ⊆ z
  384. x⊆z⇒y⊆z⇒x∪y⊆z {x = []} _ y⊆z = y⊆z
  385. x⊆z⇒y⊆z⇒x∪y⊆z {x = h ∷ t} {y = y} {z = z} x⊆z y⊆z = f where
  386. f : (h ∷ t) ∪ y ⊆ z
  387. f {e} (here e≡h) = x⊆z $ here e≡h
  388. f {e} (there e∈t∪y) = x⊆z⇒y⊆z⇒x∪y⊆z (x∪y⊆z⇒y⊆z [ h ] t x⊆z) y⊆z e∈t∪y
  389. x⊆x∖y∪y : (x y : Names) → x ⊆ x ∖ y ∪ y
  390. x⊆x∖y∪y [] _ = []⊆x _
  391. x⊆x∖y∪y (h ∷ t) y with h ∈? y
  392. x⊆x∖y∪y (h ∷ t) y | yes h∈y = x⊆z⇒y⊆z⇒x∪y⊆z (e∈l⇒[e]⊆l $ e∈x⇒e∈y∪x (t ∖ y) h∈y) (x⊆x∖y∪y t y)
  393. x⊆x∖y∪y (h ∷ t) y | no _ = x∪y⊆x̀∪ỳ ([ h ] ⊆ [ h ] ∋ ≋⇒⊆ refl) (x⊆x∖y∪y t y)
  394. e₁∈l⇒e₁∉l∖e₂⇒e₁≡e₂ : {e₁ e₂ : String} {l : Names} → e₁ ∈ l → e₁ ∉ l ∖ [ e₂ ] → e₁ ≡ e₂
  395. e₁∈l⇒e₁∉l∖e₂⇒e₁≡e₂ {e₁ = e₁} {e₂ = e₂} {l = l} e₁∈l e₁∉l∖e₂ with e₁ ≟ e₂
  396. ... | yes e₁≡e₂ = e₁≡e₂
  397. ... | no e₁≢e₂ = ⊥-elim $ p₄ e₁∈l where
  398. p₁ : e₁ ∉ [ e₂ ] ∪ (l ∖ [ e₂ ])
  399. p₁ = x≢y⇒x∉l⇒x∉y∷l e₁≢e₂ e₁∉l∖e₂
  400. p₂ : [ e₂ ] ∪ (l ∖ [ e₂ ]) ≋ (l ∖ [ e₂ ]) ∪ [ e₂ ]
  401. p₂ = ∪-sym [ e₂ ] (l ∖ [ e₂ ])
  402. p₃ : e₁ ∉ (l ∖ [ e₂ ]) ∪ [ e₂ ]
  403. p₃ = p₁ ∘ ≋⇒⊆ (∪-sym (l ∖ [ e₂ ]) [ e₂ ])
  404. p₄ : e₁ ∉ l
  405. p₄ = x⊆y⇒e∉y⇒e∉x (x⊆x∖y∪y l [ e₂ ]) p₃
  406. e∉x⇒x∩y≋x∩y∖e : {e : String} {x : Names} (y : Names) → e ∉ x → x ∩ y ≡ x ∩ (y ∖ [ e ])
  407. e∉x⇒x∩y≋x∩y∖e {x = []} _ _ = refl
  408. e∉x⇒x∩y≋x∩y∖e {e = e} {x = h ∷ t} y e∉x with h ∈? y | h ∈? (y ∖ [ e ])
  409. ... | yes h∈y | yes h∈y∖e = ≡-cong (_∷_ h) $ e∉x⇒x∩y≋x∩y∖e y $ x∉y∷l⇒x∉l e∉x
  410. ... | yes h∈y | no h∉y∖e = ⊥-elim $ e∉x e∈x where
  411. e∈x : e ∈ h ∷ t
  412. e∈x = here $ ≡-sym $ e₁∈l⇒e₁∉l∖e₂⇒e₁≡e₂ h∈y h∉y∖e
  413. ... | no h∉y | yes h∈y∖e = ⊥-elim $ h∉y $ x∖y⊆x y [ e ] h∈y∖e
  414. ... | no h∉y | no h∉y∖e = e∉x⇒x∩y≋x∩y∖e y $ x∉y∷l⇒x∉l e∉x
  415. y⊆x⇒x∩y≋y : {x y : Names} → NonRepetitive y → NonRepetitive x → y ⊆ x → x ∩ y ≋ y
  416. y⊆x⇒x∩y≋y {x = []} _ _ y⊆[] = ≡⇒≋ $ ≡-elim′ (λ y → [] ∩ y ≡ y) (≡-sym $ x⊆[]⇒x≡[] y⊆[]) refl
  417. y⊆x⇒x∩y≋y {x = h ∷ t} {y = y} nr-y (h∉t ∷ nr-t) y⊆h∷t with h ∈? y
  418. ... | yes h∈y = ≋-trans (∪-sym [ h ] (t ∩ y)) p₄ where
  419. p₁ : t ∩ y ≋ t ∩ (y ∖ [ h ])
  420. p₁ = ≡⇒≋ $ e∉x⇒x∩y≋x∩y∖e y h∉t
  421. p₂ : t ∩ (y ∖ [ h ]) ≋ y ∖ [ h ]
  422. p₂ = y⊆x⇒x∩y≋y (nr-x∖y nr-y [ h ]) nr-t $ x⊆e∷y⇒e∉x⇒x⊆y (y⊆h∷t ∘ x∖y⊆x y [ h ]) (e∉l∖[e] h y)
  423. p₃ : y ∖ [ h ] ∪ [ h ] ≋ y
  424. p₃ = e∈l⇒nr-l⇒l∖e∪e≋l h∈y nr-y
  425. p₄ : t ∩ y ∪ [ h ] ≋ y
  426. p₄ = ≋-trans (x≋x̀⇒x∪y≋x̀∪y (≋-trans p₁ p₂) [ h ]) p₃
  427. ... | no h∉y = y⊆x⇒x∩y≋y nr-y nr-t $ x⊆e∷y⇒e∉x⇒x⊆y y⊆h∷t h∉y
  428. x∪y∖n≡x∖n∪y∖n : ∀ {ℓ} {A : Set ℓ} (x y : List (Named A)) (n : Names) → (x ∪ y) ∖∖ n ≡ (x ∖∖ n) ∪ (y ∖∖ n)
  429. x∪y∖n≡x∖n∪y∖n [] _ _ = refl
  430. x∪y∖n≡x∖n∪y∖n (h ∷ t) y n with proj₁ h ∈? n
  431. ... | yes _ = x∪y∖n≡x∖n∪y∖n t y n
  432. ... | no _ = ≡-cong (_∷_ h) (x∪y∖n≡x∖n∪y∖n t y n)
  433. n-x⊆n⇒x∖n≡[] : ∀ {ℓ} {A : Set ℓ} (x : List (Named A)) (n : Names) → names x ⊆ n → x ∖∖ n ≡ []
  434. n-x⊆n⇒x∖n≡[] [] _ _ = refl
  435. n-x⊆n⇒x∖n≡[] (h ∷ t) n n-x⊆n with proj₁ h ∈? n
  436. ... | yes _ = n-x⊆n⇒x∖n≡[] t n $ x∪y⊆z⇒y⊆z [ proj₁ h ] (names t) (x≋y⊆z (≡⇒≋ $ n-x∪y [ h ] t) n-x⊆n)
  437. ... | no h∉n = ⊥-elim $ h∉n $ n-x⊆n (proj₁ h ∈ proj₁ h ∷ names t ∋ here refl)
  438. x∖x≡[] : ∀ {ℓ} {A : Set ℓ} (x : List (Named A)) → x ∖∖ names x ≡ []
  439. x∖x≡[] x = n-x⊆n⇒x∖n≡[] x (names x) (≋⇒⊆ refl)
  440. nr-[a] : ∀ {ℓ} {A : Set ℓ} {a : A} → NonRepetitive [ a ]
  441. nr-[a] = (λ()) ∷ []
  442. x≋y⇒x∖n≋y∖n : ∀ {ℓ} {A : Set ℓ} {x y : List (Named A)} → x ≋ y → (n : Names) → x ∖∖ n ≋ y ∖∖ n
  443. x≋y⇒x∖n≋y∖n refl _ = refl
  444. x≋y⇒x∖n≋y∖n {x = x} (perm l₁ x₁ x₂ l₂ p) n = ≋-trans p₀ $ ≋-trans (≡⇒≋ p₅) $ ≋-trans pg (≡⇒≋ $ ≡-sym g₅) where
  445. p₀ : x ∖∖ n ≋ (l₁ ++ x₂ ∷ x₁ ∷ l₂) ∖∖ n
  446. p₀ = x≋y⇒x∖n≋y∖n p n
  447. p₁ : (l₁ ++ [ x₂ ] ++ [ x₁ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ ([ x₂ ] ++ [ x₁ ] ++ l₂) ∖∖ n
  448. p₁ = x∪y∖n≡x∖n∪y∖n l₁ ([ x₂ ] ++ [ x₁ ] ++ l₂) n
  449. p₂ : l₁ ∖∖ n ++ ([ x₂ ] ++ [ x₁ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ [ x₂ ] ∖∖ n ++ ([ x₁ ] ++ l₂) ∖∖ n
  450. p₂ = ≡-cong (λ z → l₁ ∖∖ n ++ z) (x∪y∖n≡x∖n∪y∖n [ x₂ ] ([ x₁ ] ++ l₂) n)
  451. p₃ : l₁ ∖∖ n ++ [ x₂ ] ∖∖ n ++ ([ x₁ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ [ x₂ ] ∖∖ n ++ [ x₁ ] ∖∖ n ++ l₂ ∖∖ n
  452. p₃ = ≡-cong (λ z → l₁ ∖∖ n ++ [ x₂ ] ∖∖ n ++ z) (x∪y∖n≡x∖n∪y∖n [ x₁ ] l₂ n)
  453. p₄ : l₁ ∖∖ n ++ [ x₂ ] ∖∖ n ++ [ x₁ ] ∖∖ n ++ l₂ ∖∖ n ≡ l₁ ∖∖ n ++ ([ x₂ ] ∖∖ n ++ [ x₁ ] ∖∖ n) ++ l₂ ∖∖ n
  454. p₄ = ≡-cong (λ z → l₁ ∖∖ n ++ z) (≡-sym $ ∪-assoc ([ x₂ ] ∖∖ n) ([ x₁ ] ∖∖ n) (l₂ ∖∖ n))
  455. p₅ : (l₁ ++ [ x₂ ] ++ [ x₁ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ ([ x₂ ] ∖∖ n ++ [ x₁ ] ∖∖ n) ++ l₂ ∖∖ n
  456. p₅ = ≡-trans p₁ $ ≡-trans p₂ $ ≡-trans p₃ p₄
  457. pg : l₁ ∖∖ n ++ ([ x₂ ] ∖∖ n ++ [ x₁ ] ∖∖ n) ++ l₂ ∖∖ n ≋ l₁ ∖∖ n ++ ([ x₁ ] ∖∖ n ++ [ x₂ ] ∖∖ n) ++ l₂ ∖∖ n
  458. pg = y≋ỳ⇒x∪y≋x∪ỳ (l₁ ∖∖ n) $ x≋x̀⇒x∪y≋x̀∪y (∪-sym ([ x₂ ] ∖∖ n) ([ x₁ ] ∖∖ n)) (l₂ ∖∖ n)
  459. g₁ : (l₁ ++ [ x₁ ] ++ [ x₂ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ ([ x₁ ] ++ [ x₂ ] ++ l₂) ∖∖ n
  460. g₁ = x∪y∖n≡x∖n∪y∖n l₁ ([ x₁ ] ++ [ x₂ ] ++ l₂) n
  461. g₂ : l₁ ∖∖ n ++ ([ x₁ ] ++ [ x₂ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ [ x₁ ] ∖∖ n ++ ([ x₂ ] ++ l₂) ∖∖ n
  462. g₂ = ≡-cong (λ z → l₁ ∖∖ n ++ z) (x∪y∖n≡x∖n∪y∖n [ x₁ ] ([ x₂ ] ++ l₂) n)
  463. g₃ : l₁ ∖∖ n ++ [ x₁ ] ∖∖ n ++ ([ x₂ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ [ x₁ ] ∖∖ n ++ [ x₂ ] ∖∖ n ++ l₂ ∖∖ n
  464. g₃ = ≡-cong (λ z → l₁ ∖∖ n ++ [ x₁ ] ∖∖ n ++ z) (x∪y∖n≡x∖n∪y∖n [ x₂ ] l₂ n)
  465. g₄ : l₁ ∖∖ n ++ [ x₁ ] ∖∖ n ++ [ x₂ ] ∖∖ n ++ l₂ ∖∖ n ≡ l₁ ∖∖ n ++ ([ x₁ ] ∖∖ n ++ [ x₂ ] ∖∖ n) ++ l₂ ∖∖ n
  466. g₄ = ≡-cong (λ z → l₁ ∖∖ n ++ z) (≡-sym $ ∪-assoc ([ x₁ ] ∖∖ n) ([ x₂ ] ∖∖ n) (l₂ ∖∖ n))
  467. g₅ : (l₁ ++ [ x₁ ] ++ [ x₂ ] ++ l₂) ∖∖ n ≡ l₁ ∖∖ n ++ ([ x₁ ] ∖∖ n ++ [ x₂ ] ∖∖ n) ++ l₂ ∖∖ n
  468. g₅ = ≡-trans g₁ $ ≡-trans g₂ $ ≡-trans g₃ g₄
  469. nr-x∪y⇒e∈x⇒e∈y⇒⊥ : ∀ {ℓ} {A : Set ℓ} {e : A} {x y : List A} → NonRepetitive (x ∪ y) → e ∈ x → e ∈ y → ⊥
  470. nr-x∪y⇒e∈x⇒e∈y⇒⊥ {x = []} _ ()
  471. nr-x∪y⇒e∈x⇒e∈y⇒⊥ {x = h ∷ t} (h∉t∪y ∷ nr-t∪y) (here e≡h) = h∉t∪y ∘ e∈x⇒e∈y∪x t ∘ ≡-elim′ (λ x → x ∈ _) e≡h
  472. nr-x∪y⇒e∈x⇒e∈y⇒⊥ {x = h ∷ t} (h∉t∪y ∷ nr-t∪y) (there e∈t) = nr-x∪y⇒e∈x⇒e∈y⇒⊥ nr-t∪y e∈t
  473. nr-x∪y⇒x∖y≡x : ∀ {ℓ} {A : Set ℓ} (x y : List (Named A)) → NonRepetitiveNames (x ∪ y) → y ∖∖ names x ≡ y
  474. nr-x∪y⇒x∖y≡x x y nr-x∪y = f₀ y (≋⇒⊆ refl) where
  475. f₀ : ∀ t → t ⊆ y → t ∖∖ names x ≡ t
  476. f₀ [] _ = refl
  477. f₀ (h ∷ t) h∷t⊆y with proj₁ h ∈? names x
  478. ... | yes h∈x = ⊥-elim $ nr-x∪y⇒e∈x⇒e∈y⇒⊥ (≡-elim′ NonRepetitive (n-x∪y x y) nr-x∪y) h∈x (x⊆y⇒nx⊆ny h∷t⊆y $ here refl)
  479. ... | no _ = ≡-cong (_∷_ h) (f₀ t $ x∪y⊆z⇒y⊆z [ h ] t h∷t⊆y)
  480. t≋t₁∪t₂⇒t∖t₁≋t₂ : ∀ {ℓ} {A : Set ℓ} {t : List (Named A)} → NonRepetitiveNames t → (t₁ t₂ : List (Named A)) → t ≋ t₁ ∪ t₂ → t ∖∖ names t₁ ≋ t₂
  481. t≋t₁∪t₂⇒t∖t₁≋t₂ {t = t} nr-t t₁ t₂ t≋t₁∪t₂ = ≋-trans p₂ (≡⇒≋ $ nr-x∪y⇒x∖y≡x t₁ t₂ $ nr-x≋y (n-x≋y t≋t₁∪t₂) nr-t) where
  482. n-t₁ = names t₁
  483. p₁ : t ∖∖ n-t₁ ≋ (t₁ ∖∖ n-t₁) ∪ (t₂ ∖∖ n-t₁)
  484. p₁ = ≋-trans (x≋y⇒x∖n≋y∖n t≋t₁∪t₂ n-t₁) (≡⇒≋ $ x∪y∖n≡x∖n∪y∖n t₁ t₂ n-t₁)
  485. p₂ : t ∖∖ n-t₁ ≋ t₂ ∖∖ n-t₁
  486. p₂ = ≡-elim′ (λ x → t ∖∖ n-t₁ ≋ x ∪ (t₂ ∖∖ n-t₁)) (x∖x≡[] t₁) p₁
  487. x≋x̀⇒y≋ỳ⇒x∪y≋x̀∪ỳ : ∀ {ℓ} {A : Set ℓ} {x x̀ y ỳ : List A} → x ≋ x̀ → y ≋ ỳ → x ∪ y ≋ x̀ ∪ ỳ
  488. x≋x̀⇒y≋ỳ⇒x∪y≋x̀∪ỳ x≋x̀ (refl {y}) = x≋x̀⇒x∪y≋x̀∪y x≋x̀ y
  489. x≋x̀⇒y≋ỳ⇒x∪y≋x̀∪ỳ {x = x} {x̀ = x̀} {y = y} {ỳ = ._} x≋x̀ (perm l₁ x₁ x₂ l₂ y≋l₁x₂x₁l₂) = ≋-trans p₁ $ ≋-trans p₂ $ ≋-trans p₃ p₄ where
  490. p₁ : x ∪ y ≋ x̀ ∪ (l₁ ++ x₂ ∷ x₁ ∷ l₂)
  491. p₁ = x≋x̀⇒y≋ỳ⇒x∪y≋x̀∪ỳ x≋x̀ y≋l₁x₂x₁l₂
  492. p₂ : x̀ ∪ (l₁ ++ x₂ ∷ x₁ ∷ l₂) ≋ (x̀ ∪ l₁) ++ x₂ ∷ x₁ ∷ l₂
  493. p₂ = ≡⇒≋ $ ≡-sym $ ∪-assoc x̀ l₁ (x₂ ∷ x₁ ∷ l₂)
  494. p₃ : (x̀ ∪ l₁) ++ x₂ ∷ x₁ ∷ l₂ ≋ (x̀ ∪ l₁) ++ x₁ ∷ x₂ ∷ l₂
  495. p₃ = perm (x̀ ∪ l₁) x₁ x₂ l₂ refl
  496. p₄ : (x̀ ∪ l₁) ++ x₁ ∷ x₂ ∷ l₂ ≋ x̀ ∪ (l₁ ++ x₁ ∷ x₂ ∷ l₂)
  497. p₄ = ≡⇒≋ $ ∪-assoc x̀ l₁ (x₁ ∷ x₂ ∷ l₂)
  498. x∖y∪y≋y∖x∪x : ∀ {ℓ} {A : Set ℓ} (x y : List (Named A)) →
  499. let nx = names x
  500. ny = names y
  501. in filter-∈ x ny ≋ filter-∈ y nx → x ∖∖ ny ∪ y ≋ y ∖∖ nx ∪ x
  502. x∖y∪y≋y∖x∪x x y p₀ = ≋-trans p₁ $ ≋-trans p₂ p₃ where
  503. nx = names x
  504. ny = names y
  505. p₁ : x ∖∖ ny ∪ y ≋ x ∖∖ ny ∪ y ∖∖ nx ∪ filter-∈ y nx
  506. p₁ = ≋-trans (y≋ỳ⇒x∪y≋x∪ỳ (x ∖∖ ny) (t≋t∖n∪t∩n y nx)) (≡⇒≋ $ ≡-sym $ ∪-assoc (x ∖∖ ny) (y ∖∖ nx) (filter-∈ y nx))
  507. p₂ : (x ∖∖ ny ∪ y ∖∖ nx) ∪ filter-∈ y nx ≋ (y ∖∖ nx ∪ x ∖∖ ny) ∪ filter-∈ x ny
  508. p₂ = x≋x̀⇒y≋ỳ⇒x∪y≋x̀∪ỳ (∪-sym (x ∖∖ ny) (y ∖∖ nx)) (≋-sym p₀)
  509. p₃ : y ∖∖ nx ∪ x ∖∖ ny ∪ filter-∈ x ny ≋ y ∖∖ nx ∪ x
  510. p₃ = ≋-trans (≡⇒≋ $ ∪-assoc (y ∖∖ nx) (x ∖∖ ny) (filter-∈ x ny)) (≋-sym $ y≋ỳ⇒x∪y≋x∪ỳ (y ∖∖ nx) (t≋t∖n∪t∩n x ny))
  511. names⇒named : Names → List (Named Unit)
  512. names⇒named = mapL (λ x → x , unit)
  513. nn-x∪y≡nn-x∪nn-y : ∀ x y → names⇒named (x ∪ y) ≡ names⇒named x ∪ names⇒named y
  514. nn-x∪y≡nn-x∪nn-y [] y = refl
  515. nn-x∪y≡nn-x∪nn-y (h ∷ t) y = ≡-cong (_∷_ _) (nn-x∪y≡nn-x∪nn-y t y)
  516. nn-x≡x : ∀ x → names (names⇒named x) ≡ x
  517. nn-x≡x [] = refl
  518. nn-x≡x (h ∷ t) = ≡-cong (_∷_ h) (nn-x≡x t)
  519. x≋y⇒nn-x≋nn-y : ∀ {x y} → x ≋ y → names⇒named x ≋ names⇒named y
  520. x≋y⇒nn-x≋nn-y refl = refl
  521. x≋y⇒nn-x≋nn-y {x = x} {y = ._} (perm l₁ x₁ x₂ l₂ x≋l₁x₂x₁l₂) = ≋-trans p₁ p₂ where
  522. p₁ : names⇒named x ≋ names⇒named l₁ ++ (x₂ , unit) ∷ (x₁ , unit) ∷ names⇒named l₂
  523. p₁ = ≋-trans (x≋y⇒nn-x≋nn-y x≋l₁x₂x₁l₂) (≡⇒≋ $ nn-x∪y≡nn-x∪nn-y l₁ (x₂ ∷ x₁ ∷ l₂))
  524. p₂ : names⇒named l₁ ++ (x₂ , unit) ∷ (x₁ , unit) ∷ names⇒named l₂ ≋ names⇒named (l₁ ++ x₁ ∷ x₂ ∷ l₂)
  525. p₂ = ≋-trans (perm (names⇒named l₁) (x₁ , unit) (x₂ , unit) (names⇒named l₂) refl) (≡⇒≋ $ ≡-sym $ nn-x∪y≡nn-x∪nn-y l₁ (x₁ ∷ x₂ ∷ l₂))
  526. x≋y∪z⇒x∖y≋z : {x : Names} → NonRepetitive x → (y z : Names) → x ≋ y ∪ z → x ∖ y ≋ z
  527. x≋y∪z⇒x∖y≋z {x} nr-x y z x≋y∪z = ≋-trans (≡⇒≋ $ ≡-sym $ ≡-trans p₂ p₃) (≋-trans p₁ $ ≡⇒≋ $ nn-x≡x z) where
  528. ux = names⇒named x
  529. uy = names⇒named y
  530. uz = names⇒named z
  531. ux≋uy∪uz : ux ≋ uy ∪ uz
  532. ux≋uy∪uz = ≋-trans (x≋y⇒nn-x≋nn-y x≋y∪z) (≡⇒≋ $ nn-x∪y≡nn-x∪nn-y y z)
  533. p₁ : names (ux ∖∖ names uy) ≋ names uz
  534. p₁ = n-x≋y $ t≋t₁∪t₂⇒t∖t₁≋t₂ (≡-elim′ NonRepetitive (≡-sym $ nn-x≡x x) nr-x) uy uz ux≋uy∪uz
  535. p₂ : names (ux ∖∖ names uy) ≡ names ux ∖ names uy
  536. p₂ = n-x∖y ux (names uy)
  537. p₃ : names ux ∖ names uy ≡ x ∖ y
  538. p₃ = ≡-cong₂ _∖_ (nn-x≡x x) (nn-x≡x y)
  539. h∈x⇒∃t,x≋h∷t : ∀ {ℓ} {A : Set ℓ} {h : A} {x : List A} → h ∈ x → NonRepetitive x → Σ[ t ∈ List A ] x ≋ h ∷ t
  540. h∈x⇒∃t,x≋h∷t {h = h} .{x = x ∷ xs} (here {x = x} {xs = xs} h≡x) (x∉xs ∷ nr-xs) = xs , ≡⇒≋ (≡-cong (λ z → z ∷ xs) (≡-sym h≡x))
  541. h∈x⇒∃t,x≋h∷t {h = h} .{x = x ∷ xs} (there {x = x} {xs = xs} h∈xs) (x∉xs ∷ nr-xs) =
  542. let t , xs≋h∷t = ((Σ[ t ∈ List _ ] xs ≋ h ∷ t) ∋ h∈x⇒∃t,x≋h∷t h∈xs nr-xs)
  543. p₁ : x ∷ xs ≋ x ∷ h ∷ t
  544. p₁ = y≋ỳ⇒x∪y≋x∪ỳ [ x ] xs≋h∷t
  545. p₂ : x ∷ h ∷ t ≋ h ∷ x ∷ t
  546. p₂ = ≋-del-ins-l [] [ x ] h t
  547. in x ∷ t , ≋-trans p₁ p₂
  548. nr-x∪y⇒nr-y : ∀ {ℓ} {A : Set ℓ} (x y : List A) → NonRepetitive (x ∪ y) → NonRepetitive y
  549. nr-x∪y⇒nr-y [] _ nr-y = nr-y
  550. nr-x∪y⇒nr-y (h ∷ t) y (_ ∷ nr-t∪y) = nr-x∪y⇒nr-y t y nr-t∪y
  551. x⊆y⇒∃x̀,y≋x∪x̀ : ∀ {ℓ} {A : Set ℓ} {x y : List A} → NonRepetitive x → NonRepetitive y → x ⊆ y → Σ[ x̀ ∈ List A ] y ≋ x ∪ x̀
  552. x⊆y⇒∃x̀,y≋x∪x̀ {x = []} {y = y} _ _ _ = y , refl
  553. x⊆y⇒∃x̀,y≋x∪x̀ {x = h ∷ t} {y = y} (h∉t ∷ nr-t) nr-y h∷t⊆y =
  554. let ỳ , y≋h∷ỳ = h∈x⇒∃t,x≋h∷t (h∷t⊆y (h ∈ h ∷ t ∋ here refl)) nr-y
  555. p₁ : t ⊆ ỳ
  556. p₁ = x⊆e∷y⇒e∉x⇒x⊆y (x⊆y≋z (x∪y⊆z⇒y⊆z [ h ] t h∷t⊆y) y≋h∷ỳ) h∉t
  557. nr-ỳ = nr-x∪y⇒nr-y [ h ] ỳ $ nr-x≋y y≋h∷ỳ nr-y
  558. t̀ , ỳ≋t∪t̀ = x⊆y⇒∃x̀,y≋x∪x̀ nr-t nr-ỳ p₁
  559. p₂ : h ∷ ỳ ≋ (h ∷ t) ∪ t̀
  560. p₂ = y≋ỳ⇒x∪y≋x∪ỳ [ h ] ỳ≋t∪t̀
  561. in t̀ , ≋-trans y≋h∷ỳ p₂
  562. t₁⊆t₂⇒t₂≋t₁∪t₂∖nt₁ : ∀ {ℓ} {A : Set ℓ} {t₁ t₂ : List (Named A)} → NonRepetitiveNames t₁ → NonRepetitiveNames t₂ → t₁ ⊆ t₂ → t₂ ≋ t₁ ∪ t₂ ∖∖ names t₁
  563. t₁⊆t₂⇒t₂≋t₁∪t₂∖nt₁ {t₁ = t₁} {t₂ = t₂} nr-nt₁ nr-nt₂ t₁⊆t₂ = ≋-trans p₁ $ y≋ỳ⇒x∪y≋x∪ỳ t₁ p₆ where
  564. n-t₁ = names t₁
  565. nr-t₁ = nr-names⇒nr nr-nt₁
  566. nr-t₂ = nr-names⇒nr nr-nt₂
  567. t̀₁,t₁∪t̀₁≋t₂ = x⊆y⇒∃x̀,y≋x∪x̀ nr-t₁ nr-t₂ t₁⊆t₂
  568. t̀₁ = proj₁ t̀₁,t₁∪t̀₁≋t₂
  569. p₁ : t₂ ≋ t₁ ∪ t̀₁
  570. p₁ = proj₂ t̀₁,t₁∪t̀₁≋t₂
  571. p₂ : t₂ ∖∖ n-t₁ ≋ (t₁ ∪ t̀₁) ∖∖ n-t₁
  572. p₂ = x≋y⇒x∖n≋y∖n p₁ n-t₁
  573. p₃ : (t₁ ∪ t̀₁) ∖∖ n-t₁ ≡ t₁ ∖∖ n-t₁ ∪ t̀₁ ∖∖ n-t₁
  574. p₃ = x∪y∖n≡x∖n∪y∖n t₁ t̀₁ n-t₁
  575. p₄ : t₁ ∖∖ n-t₁ ∪ t̀₁ ∖∖ n-t₁ ≡ t̀₁ ∖∖ n-t₁
  576. p₄ = ≡-cong (λ x → x ∪ t̀₁ ∖∖ n-t₁) (x∖x≡[] t₁)
  577. p₅ : t̀₁ ∖∖ n-t₁ ≡ t̀₁
  578. p₅ = nr-x∪y⇒x∖y≡x t₁ t̀₁ $ nr-x≋y (n-x≋y p₁) nr-nt₂
  579. p₆ : t̀₁ ≋ t₂ ∖∖ n-t₁
  580. p₆ = ≋-sym $ ≋-trans p₂ $ ≡⇒≋ $ ≡-trans p₃ $ ≡-trans p₄ p₅
  581. t₁⊆t₂⇒t₁≋f∈-t₂-nt₁ : ∀ {ℓ} {A : Set ℓ} {t₁ t₂ : List (Named A)} → NonRepetitiveNames t₁ → NonRepetitiveNames t₂ → t₁ ⊆ t₂ → t₁ ≋ filter-∈ t₂ (names t₁)
  582. t₁⊆t₂⇒t₁≋f∈-t₂-nt₁ {t₁ = t₁} {t₂ = t₂} nr-t₁ nr-t₂ t₁⊆t₂ = ≋-trans (≋-sym p₃) p₄ where
  583. nt₁ = names t₁
  584. p₁ : t₂ ≋ t₂ ∖∖ nt₁ ∪ t₁
  585. p₁ = ≋-trans (t₁⊆t₂⇒t₂≋t₁∪t₂∖nt₁ nr-t₁ nr-t₂ t₁⊆t₂) (∪-sym t₁ (t₂ ∖∖ nt₁))
  586. p₂ : t₂ ≋ t₂ ∖∖ nt₁ ∪ filter-∈ t₂ nt₁
  587. p₂ = t≋t∖n∪t∩n t₂ nt₁
  588. p₃ : t₂ ∖∖ names (t₂ ∖∖ nt₁) ≋ t₁
  589. p₃ = t≋t₁∪t₂⇒t∖t₁≋t₂ nr-t₂ (t₂ ∖∖ nt₁) t₁ p₁
  590. p₄ : t₂ ∖∖ names (t₂ ∖∖ nt₁) ≋ filter-∈ t₂ nt₁
  591. p₄ = t≋t₁∪t₂⇒t∖t₁≋t₂ nr-t₂ (t₂ ∖∖ nt₁) (filter-∈ t₂ nt₁) p₂
  592. -- /lemmas
  593. -- assertions
  594. infix 4 _s-≢!_
  595. _s-≢!_ : String → String → Set
  596. x s-≢! y with x ≟ y
  597. ... | yes _ = ⊥
  598. ... | no _ = ⊤
  599. s-≢!⇒≢? : ∀ {x y} → x s-≢! y → x ≢ y
  600. s-≢!⇒≢? {x} {y} x≢!y with x ≟ y
  601. s-≢!⇒≢? () | yes _
  602. s-≢!⇒≢? _ | no p = p
  603. s-NonRepetitive? : (n : Names) → Dec (NonRepetitive n)
  604. s-NonRepetitive? [] = yes []
  605. s-NonRepetitive? (h ∷ t) with h ∈? t | s-NonRepetitive? t
  606. ... | no h∉t | yes nr-t = yes (h∉t ∷ nr-t)
  607. ... | yes h∈t | _ = no f where
  608. f : NonRepetitive (_ ∷ _) → _
  609. f (h∉t ∷ _) = h∉t h∈t
  610. ... | _ | no ¬nr-t = no f where
  611. f : NonRepetitive (_ ∷ _) → _
  612. f (_ ∷ nr-t) = ¬nr-t nr-t
  613. s-NonRepetitive! : Names → Set
  614. s-NonRepetitive! n with s-NonRepetitive? n
  615. ... | yes _ = ⊤
  616. ... | no _ = ⊥
  617. NonRepetitiveNames! : ∀ {ℓ} {A : Set ℓ} → List (Named A) → Set
  618. NonRepetitiveNames! = s-NonRepetitive! ∘ names
  619. s-nr!⇒nr : {n : Names} → s-NonRepetitive! n → NonRepetitive n
  620. s-nr!⇒nr {n = n} _ with s-NonRepetitive? n
  621. s-nr!⇒nr _ | yes p = p
  622. s-nr!⇒nr () | no _
  623. -- /assertions