Transformer.agda 12 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196
  1. module Issue784.Transformer where
  2. open import Data.List using (List; []; _∷_; _++_; [_]; filter) renaming (map to mapL)
  3. import Level
  4. open import Issue784.Values
  5. open import Issue784.Context
  6. Transformer : ∀ {ℓ} → NonRepetitiveTypes ℓ → NonRepetitiveTypes ℓ → Set (Level.suc ℓ)
  7. Transformer {ℓ} (t-in , nr-in) (t-out , nr-out) =
  8. (v : Context ℓ) → NonRepetitiveContext v → t-in ⊆ signature v → NonRepetitive (ctxnames v ∖ names t-in ∪ names t-out) →
  9. Σ[ w ∈ Context ℓ ] signature w ≋ signature v ∖∖ names t-in ∪ t-out
  10. pipe : ∀ {ℓ} {t-in₁ t-out₁ t-in₂ t-out₂ : Types ℓ} {nr-in₁ nr-out₁ nr-in₂ nr-out₂} →
  11. Transformer (t-in₁ , nr-in₁) (t-out₁ , nr-out₁) →
  12. Transformer (t-in₂ , nr-in₂) (t-out₂ , nr-out₂) →
  13. let n-out₁ = names t-out₁
  14. n-in₂ = names t-in₂
  15. t-in = t-in₁ ∪ (t-in₂ ∖∖ n-out₁)
  16. t-out = t-out₁ ∖∖ n-in₂ ∪ t-out₂
  17. in
  18. filter-∈ t-out₁ n-in₂ ≋ filter-∈ t-in₂ n-out₁ →
  19. (nr-in : NonRepetitiveNames t-in) (nr-out : NonRepetitiveNames t-out) →
  20. Transformer (t-in , nr-in) (t-out , nr-out)
  21. pipe {ℓ} {t-in₁} {t-out₁} {t-in₂} {t-out₂} {nr-in₁} {nr-out₁} {nr-in₂} {nr-out₂} tr₁ tr₂ pr-t nr-in nr-out = tr where
  22. n-in₁ = names t-in₁
  23. n-out₁ = names t-out₁
  24. n-in₂ = names t-in₂
  25. n-out₂ = names t-out₂
  26. t-in = t-in₁ ∪ (t-in₂ ∖∖ n-out₁)
  27. t-out = (t-out₁ ∖∖ n-in₂) ∪ t-out₂
  28. tr : Transformer (t-in , nr-in) (t-out , nr-out)
  29. tr ctx nr-v t-ì⊆v nr-ò = context w , w≋out where
  30. v = Context.get ctx
  31. n-in = names t-in
  32. v̀ = filter-∈ v n-in
  33. nr-v̀ : NonRepetitiveNames v̀
  34. nr-v̀ = nr-x⇒nr-x∩y nr-v n-in
  35. v̀≋i : types v̀ ≋ t-in
  36. v̀≋i = ≋-sym $ ≋-trans p₁ p₂ where
  37. p₁ : t-in ≋ filter-∈ (types v) n-in
  38. p₁ = t₁⊆t₂⇒t₁≋f∈-t₂-nt₁ nr-in (≡-elim′ NonRepetitive (≡-sym $ n-types v) nr-v) t-ì⊆v
  39. p₂ : filter-∈ (types v) n-in ≋ types v̀
  40. p₂ = ≡⇒≋ $ ≡-sym $ t-filter-∈ v n-in
  41. -- transformer₁
  42. i₁⊆v̀ : t-in₁ ⊆ types v̀
  43. i₁⊆v̀ = x⊆y≋z (x⊆x∪y t-in₁ (t-in₂ ∖∖ n-out₁)) (≋-sym v̀≋i)
  44. v̀∖i₁∪o₁≋i₂∖o₁∪o₁ : types v̀ ∖∖ n-in₁ ∪ t-out₁ ≋ t-in₂ ∖∖ n-out₁ ∪ t-out₁
  45. v̀∖i₁∪o₁≋i₂∖o₁∪o₁ = x≋x̀⇒x∪y≋x̀∪y p₂ t-out₁ where
  46. p₁ : NonRepetitiveNames (types v̀)
  47. p₁ = nr-x≋y (≡⇒≋ $ ≡-sym $ n-types v̀) nr-v̀
  48. p₂ : types v̀ ∖∖ n-in₁ ≋ t-in₂ ∖∖ n-out₁
  49. p₂ = t≋t₁∪t₂⇒t∖t₁≋t₂ p₁ t-in₁ (t-in₂ ∖∖ n-out₁) v̀≋i
  50. n-v̀∖i₁∪o₁≋i₂∖o₁∪o₁ : names v̀ ∖ n-in₁ ∪ n-out₁ ≋ n-in₂ ∖ n-out₁ ∪ n-out₁
  51. n-v̀∖i₁∪o₁≋i₂∖o₁∪o₁ = x≋x̀⇒x∪y≋x̀∪y p₃ n-out₁ where
  52. p₁ : names v̀ ≋ n-in
  53. p₁ = ≋-trans (≡⇒≋ $ n-filter-∈ v n-in) (y⊆x⇒x∩y≋y nr-in nr-v (≡-elim′ (λ x → n-in ⊆ x) (n-types v) (x⊆y⇒nx⊆ny t-ì⊆v)))
  54. p₂ : n-in ≋ n-in₁ ∪ (n-in₂ ∖ n-out₁)
  55. p₂ = ≡⇒≋ $ ≡-trans (n-x∪y t-in₁ $ t-in₂ ∖∖ n-out₁) (≡-cong (λ x → n-in₁ ∪ x) (n-x∖y t-in₂ n-out₁))
  56. p₃ : names v̀ ∖ n-in₁ ≋ n-in₂ ∖ n-out₁
  57. p₃ = x≋y∪z⇒x∖y≋z (nr-x≋y (≋-sym p₁) nr-in) n-in₁ (n-in₂ ∖ n-out₁) (≋-trans p₁ p₂)
  58. nr-v̀∖i₁∪o₁ : NonRepetitive (names v̀ ∖ n-in₁ ∪ n-out₁)
  59. nr-v̀∖i₁∪o₁ = nr-x≋y (≋-sym n-v̀∖i₁∪o₁≋i₂∖o₁∪o₁) p₁ where
  60. p₁ : NonRepetitive (n-in₂ ∖ n-out₁ ∪ n-out₁)
  61. p₁ = nr-x∖y∪y nr-in₂ nr-out₁
  62. w-all₁ : Σ[ w ∈ Context ℓ ] signature w ≋ signature (context v̀) ∖∖ names t-in₁ ∪ t-out₁
  63. w-all₁ = tr₁ (context v̀) nr-v̀ i₁⊆v̀ nr-v̀∖i₁∪o₁
  64. w₁ = Context.get $ proj₁ w-all₁
  65. w₁≋v̀∖i₁∪o₁ = proj₂ w-all₁
  66. -- transformer₂
  67. n-w₁≋v̀∖i₁∪o₁ : names w₁ ≋ names v̀ ∖ names t-in₁ ∪ names t-out₁
  68. n-w₁≋v̀∖i₁∪o₁ = ≡-elim p₆ p₅ where
  69. p₁ : names (types w₁) ≋ names (types v̀ ∖∖ names t-in₁ ∪ t-out₁)
  70. p₁ = n-x≋y w₁≋v̀∖i₁∪o₁
  71. p₂ : names (types v̀ ∖∖ names t-in₁ ∪ t-out₁) ≡ names (types v̀ ∖∖ names t-in₁) ∪ names t-out₁
  72. p₂ = n-x∪y (types v̀ ∖∖ names t-in₁) t-out₁
  73. p₃ : names (types v̀ ∖∖ names t-in₁) ≡ names (types v̀) ∖ names t-in₁
  74. p₃ = n-x∖y (types v̀) (names t-in₁)
  75. p₄ : names (types w₁) ≋ names (types v̀ ∖∖ names t-in₁) ∪ names t-out₁
  76. p₄ = ≡-elim′ (λ x → names (types w₁) ≋ x) p₂ p₁
  77. p₅ : names (types w₁) ≋ names (types v̀) ∖ names t-in₁ ∪ names t-out₁
  78. p₅ = ≡-elim′ (λ x → names (types w₁) ≋ x ∪ names t-out₁) p₃ p₄
  79. p₆ : (names (types w₁) ≋ names (types v̀) ∖ names t-in₁ ∪ names t-out₁) ≡
  80. (names w₁ ≋ names v̀ ∖ names t-in₁ ∪ names t-out₁)
  81. p₆ = ≡-cong₂ (λ x y → x ≋ y ∖ names t-in₁ ∪ names t-out₁) (n-types w₁) (n-types v̀)
  82. nr-w₁ : NonRepetitiveNames w₁
  83. nr-w₁ = nr-x≋y (≋-sym n-w₁≋v̀∖i₁∪o₁) nr-v̀∖i₁∪o₁
  84. i₂⊆w₁ : t-in₂ ⊆ types w₁
  85. i₂⊆w₁ = x⊆y≋z (x≋y⊆z (≋-sym p₁) p₄) (≋-sym w₁≋v̀∖i₁∪o₁) where
  86. p₁ : t-in₂ ≋ (t-in₂ ∖∖ n-out₁) ∪ filter-∈ t-in₂ n-out₁
  87. p₁ = t≋t∖n∪t∩n t-in₂ n-out₁
  88. p₂ : filter-∈ t-in₂ n-out₁ ⊆ t-out₁
  89. p₂ = x≋y⊆z pr-t $ x∩y⊆x t-out₁ n-in₂
  90. p₃ : types v̀ ∖∖ n-in₁ ≋ t-in₂ ∖∖ n-out₁
  91. p₃ = t≋t₁∪t₂⇒t∖t₁≋t₂ (≡-elim′ NonRepetitive (≡-sym $ n-types v̀) nr-v̀) t-in₁ (t-in₂ ∖∖ n-out₁) v̀≋i
  92. p₄ : (t-in₂ ∖∖ n-out₁) ∪ (filter-∈ t-in₂ n-out₁) ⊆ (types v̀ ∖∖ n-in₁) ∪ t-out₁
  93. p₄ = x∪y⊆x̀∪ỳ (≋⇒⊆ $ ≋-sym p₃) p₂
  94. w₁∖i₂∪o₂≋out : types w₁ ∖∖ n-in₂ ∪ t-out₂ ≋ t-out
  95. w₁∖i₂∪o₂≋out = x≋x̀⇒x∪y≋x̀∪y p₄ t-out₂ where
  96. p₁ : t-in₂ ∖∖ n-out₁ ∪ t-out₁ ≋ t-out₁ ∖∖ n-in₂ ∪ t-in₂
  97. p₁ = ≋-sym $ x∖y∪y≋y∖x∪x t-out₁ t-in₂ pr-t
  98. p₂ : (t-out₁ ∖∖ n-in₂) ∪ t-in₂ ≋ t-in₂ ∪ (t-out₁ ∖∖ n-in₂)
  99. p₂ = ∪-sym (t-out₁ ∖∖ n-in₂) t-in₂
  100. p₃ : types w₁ ≋ t-in₂ ∪ (t-out₁ ∖∖ n-in₂)
  101. p₃ = ≋-trans w₁≋v̀∖i₁∪o₁ $ ≋-trans v̀∖i₁∪o₁≋i₂∖o₁∪o₁ $ ≋-trans p₁ p₂
  102. p₄ : types w₁ ∖∖ n-in₂ ≋ t-out₁ ∖∖ n-in₂
  103. p₄ = t≋t₁∪t₂⇒t∖t₁≋t₂ (nr-x≋y (≡⇒≋ $ ≡-sym $ n-types w₁) nr-w₁) t-in₂ (t-out₁ ∖∖ n-in₂) p₃
  104. nr-w₁∖i₂∪o₂ : NonRepetitive (names w₁ ∖ n-in₂ ∪ n-out₂)
  105. nr-w₁∖i₂∪o₂ = ≡-elim′ NonRepetitive p₄ p₅ where
  106. p₁ : names (types w₁ ∖∖ n-in₂ ∪ t-out₂) ≡ names (types w₁ ∖∖ n-in₂) ∪ n-out₂
  107. p₁ = n-x∪y (types w₁ ∖∖ n-in₂) t-out₂
  108. p₂ : names (types w₁ ∖∖ n-in₂) ∪ n-out₂ ≡ (names (types w₁) ∖ n-in₂) ∪ n-out₂
  109. p₂ = ≡-cong (λ x → x ∪ names t-out₂) (n-x∖y (types w₁) n-in₂)
  110. p₃ : (names (types w₁) ∖ n-in₂) ∪ n-out₂ ≡ names w₁ ∖ n-in₂ ∪ n-out₂
  111. p₃ = ≡-cong (λ x → x ∖ n-in₂ ∪ n-out₂) (n-types w₁)
  112. p₄ : names (types w₁ ∖∖ n-in₂ ∪ t-out₂) ≡ names w₁ ∖ n-in₂ ∪ n-out₂
  113. p₄ = ≡-trans p₁ $ ≡-trans p₂ p₃
  114. p₅ : NonRepetitiveNames (types w₁ ∖∖ n-in₂ ∪ t-out₂)
  115. p₅ = nr-x≋y (≋-sym $ n-x≋y w₁∖i₂∪o₂≋out) nr-out
  116. w-all₂ : Σ[ w ∈ Context ℓ ] signature w ≋ signature (context w₁) ∖∖ names t-in₂ ∪ t-out₂
  117. w-all₂ = tr₂ (context w₁) nr-w₁ i₂⊆w₁ nr-w₁∖i₂∪o₂
  118. w₂ = Context.get $ proj₁ w-all₂
  119. w₂≋out : types w₂ ≋ t-out
  120. w₂≋out = ≋-trans (proj₂ w-all₂) w₁∖i₂∪o₂≋out
  121. w = Values ℓ ∋ v ∖∖ n-in ∪ w₂
  122. w≋out : types w ≋ types v ∖∖ n-in ∪ t-out
  123. w≋out = ≋-trans (≡⇒≋ p₁) p₂ where
  124. p₁ : types (v ∖∖ n-in ∪ w₂) ≡ types v ∖∖ n-in ∪ types w₂
  125. p₁ = ≡-trans (t-x∪y (v ∖∖ n-in) w₂) (≡-cong (λ x → x ∪ types w₂) (t-x∖y v n-in))
  126. p₂ : types v ∖∖ n-in ∪ types w₂ ≋ types v ∖∖ n-in ∪ t-out
  127. p₂ = y≋ỳ⇒x∪y≋x∪ỳ (types v ∖∖ n-in) w₂≋out
  128. Transformer! : ∀ {ℓ} (t-in : Types ℓ) (t-out : Types ℓ) {nr!-in : NonRepetitiveNames! t-in} {nr!-out : NonRepetitiveNames! t-out} → Set (Level.suc ℓ)
  129. Transformer! t-in t-out {nr!-in = nr!-in} {nr!-out} = Transformer (t-in , s-nr!⇒nr nr!-in) (t-out , s-nr!⇒nr nr!-out)
  130. infix 1 _:=_
  131. _:=_ : ∀ {ℓ} {A : String → Set ℓ} (n : String) → ((n : String) → A n) → A n
  132. n := f = f n
  133. infixl 0 _⇒⟦_⟧⇒_
  134. _⇒⟦_⟧⇒_ : ∀ {ℓ} {t-in₁ t-out₁ t-in₂ t-out₂ : Types ℓ}
  135. {nr!-in₁ : NonRepetitiveNames! t-in₁}
  136. {nr!-out₁ : NonRepetitiveNames! t-out₁}
  137. {nr!-in₂ : NonRepetitiveNames! t-in₂}
  138. {nr!-out₂ : NonRepetitiveNames! t-out₂} →
  139. Transformer (t-in₁ , s-nr!⇒nr nr!-in₁) (t-out₁ , s-nr!⇒nr nr!-out₁) →
  140. let n-out₁ = names t-out₁
  141. n-in₂ = names t-in₂
  142. t-in = t-in₁ ∪ (t-in₂ ∖∖ n-out₁)
  143. t-out = t-out₁ ∖∖ n-in₂ ∪ t-out₂
  144. in
  145. filter-∈ t-out₁ n-in₂ ≋ filter-∈ t-in₂ n-out₁ →
  146. {nr!-in : NonRepetitiveNames! t-in} →
  147. {nr!-out : NonRepetitiveNames! t-out} →
  148. Transformer (t-in₂ , s-nr!⇒nr nr!-in₂) (t-out₂ , s-nr!⇒nr nr!-out₂) →
  149. Transformer (t-in , s-nr!⇒nr nr!-in) (t-out , s-nr!⇒nr nr!-out)
  150. _⇒⟦_⟧⇒_ {nr!-in₁ = nr!-in₁} {nr!-out₁ = nr!-out₁} {nr!-in₂ = nr!-in₂} {nr!-out₂ = nr!-out₂}
  151. tr₁ f≋f {nr!-in = nr!-in} {nr!-out = nr!-out} tr₂ =
  152. pipe {nr-in₁ = s-nr!⇒nr nr!-in₁} {nr-out₁ = s-nr!⇒nr nr!-out₁} {nr-in₂ = s-nr!⇒nr nr!-in₂} {nr-out₂ = s-nr!⇒nr nr!-out₂}
  153. tr₁ tr₂ f≋f (s-nr!⇒nr nr!-in) (s-nr!⇒nr nr!-out)
  154. record Pure {ℓ} (A : Set ℓ) : Set ℓ where
  155. constructor pure
  156. field get : A
  157. record Unique {ℓ} (A : Set ℓ) : Set ℓ where
  158. constructor unique
  159. field get : A
  160. extract : ∀ {ℓ} {n : String} {A : Set ℓ} → let t = [ (n , Pure A) ] in {nr!-t : NonRepetitiveNames! t} → Transformer ([] , []) (t , s-nr!⇒nr nr!-t) → A
  161. extract {n = n} {A = A} {nr!-t = nr!-t} tr =
  162. let e = n , Pure A
  163. v , t-v≋t = (Σ[ v ∈ Context _ ] signature v ≋ [ e ]) ∋ tr (context []) [] (≋⇒⊆ refl) (s-nr!⇒nr nr!-t)
  164. in Pure.get ∘ getBySignature $ a∈x⇒x≋y⇒a∈y (e ∈ [ e ] ∋ here refl) (≋-sym t-v≋t)