Skip to content

Commit

Permalink
Merge pull request #75 from VariantSync/2CC-indifferent-elimination
Browse files Browse the repository at this point in the history
2CC indifferent elimination
  • Loading branch information
pmbittner authored Sep 26, 2024
2 parents 7aa169b + e14b637 commit 81a5256
Show file tree
Hide file tree
Showing 2 changed files with 95 additions and 1 deletion.
6 changes: 5 additions & 1 deletion src/Vatras/Test/Experiments/ADT-to-TikZ-Forest.agda
Original file line number Diff line number Diff line change
Expand Up @@ -13,6 +13,7 @@ open import Vatras.Lang.CCC as CCC using (CCC)
open import Vatras.Lang.2CC using (2CC)
open import Vatras.Lang.ADT
open import Vatras.Translation.LanguageMap
open import Vatras.Translation.Lang.2CC.Idempotence String String._≟_ using (eliminate-idempotent-choices)

open import Vatras.Test.Experiment
open import Vatras.Show.Lines
Expand Down Expand Up @@ -50,8 +51,11 @@ CCC-to-ADT ccc = adt
bcc : STR2CC
bcc = proj₁ (2CC≽CCC ccc)

bcc' : STR2CC
bcc' = eliminate-idempotent-choices bcc

adt : STRADT
adt = proj₁ (ADT≽2CC bcc)
adt = proj₁ (ADT≽2CC bcc')

tikz-export-experiment : Experiment STRCCC
getName tikz-export-experiment = "Tikz-Export"
Expand Down
90 changes: 90 additions & 0 deletions src/Vatras/Translation/Lang/2CC/Idempotence.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,90 @@
open import Vatras.Framework.Definitions using (𝔸; 𝔽)
open import Relation.Binary.Definitions using (DecidableEquality)

module Vatras.Translation.Lang.2CC.Idempotence (Dimension : 𝔽) (_==_ : DecidableEquality Dimension) where

import Data.List as List
import Data.List.Properties as List
open import Data.Bool using (true; false; if_then_else_)
import Data.Bool.Properties as Bool
open import Data.Product using (_,_)
open import Function using (id)
open import Size using (Size; ∞)

open import Relation.Nullary using (Dec; yes; no)
open import Relation.Binary.PropositionalEquality as Eq using (refl; _≗_)
open Eq.≡-Reasoning

open import Vatras.Data.EqIndexedSet using (≗→≅[]; ≅[]-sym)
open import Vatras.Framework.Compiler using (LanguageCompiler)
open import Vatras.Framework.Variants as V using (Rose)
open import Vatras.Lang.2CC

_≟_ : {i : Size} {A : 𝔸} DecidableEquality (2CC Dimension i A)
_≟_ {A = _ , _≟ₐ_} (a₁ -< cs₁ >-) (a₂ -< cs₂ >-) with a₁ ≟ₐ a₂ | List.≡-dec _≟_ cs₁ cs₂
(a₁ -< cs₁ >-) ≟ (a₂ -< cs₂ >-) | yes a₁≡a₂ | yes cs₁≡cs₂ = yes (Eq.cong₂ _-<_>- a₁≡a₂ cs₁≡cs₂)
(a₁ -< cs₁ >-) ≟ (a₂ -< cs₂ >-) | yes a₁≡a₂ | no cs₁≢cs₂ = no λ where refl cs₁≢cs₂ refl
(a₁ -< cs₁ >-) ≟ (a₂ -< cs₂ >-) | no a₁≢a₂ | _ = no λ where refl a₁≢a₂ refl
(a₁ -< cs₁ >-) ≟ (D₂ ⟨ l₂ , r₂ ⟩) = no λ where ()
(D₁ ⟨ l₁ , r₁ ⟩) ≟ (a₂ -< cs₂ >-) = no λ where ()
(D₁ ⟨ l₁ , r₁ ⟩) ≟ (D₂ ⟨ l₂ , r₂ ⟩) with D₁ == D₂ | l₁ ≟ l₂ | r₁ ≟ r₂
(D₁ ⟨ l₁ , r₁ ⟩) ≟ (D₂ ⟨ l₂ , r₂ ⟩) | yes D₁≡d₂ | yes l₁≡l₂ | yes r₁≡r₂ = yes (Eq.cong₂ (λ f r f r) (Eq.cong₂ _⟨_,_⟩ D₁≡d₂ l₁≡l₂) r₁≡r₂)
(D₁ ⟨ l₁ , r₁ ⟩) ≟ (D₂ ⟨ l₂ , r₂ ⟩) | yes D₁≡d₂ | yes l₁≡l₂ | no r₁≢r₂ = no λ where refl r₁≢r₂ refl
(D₁ ⟨ l₁ , r₁ ⟩) ≟ (D₂ ⟨ l₂ , r₂ ⟩) | yes D₁≡d₂ | no l₁≢l₂ | _ = no λ where refl l₁≢l₂ refl
(D₁ ⟨ l₁ , r₁ ⟩) ≟ (D₂ ⟨ l₂ , r₂ ⟩) | no D₁≢d₂ | _ | _ = no λ where refl D₁≢d₂ refl

eliminate-idempotent-choices : {i : Size} {A : 𝔸} 2CC Dimension i A 2CC Dimension ∞ A
eliminate-idempotent-choices (a -< cs >-) = a -< List.map eliminate-idempotent-choices cs >-
eliminate-idempotent-choices (D ⟨ l , r ⟩) with eliminate-idempotent-choices l ≟ eliminate-idempotent-choices r
eliminate-idempotent-choices (D ⟨ l , r ⟩) | yes l≡r = eliminate-idempotent-choices l
eliminate-idempotent-choices (D ⟨ l , r ⟩) | no l≢r = D ⟨ eliminate-idempotent-choices l , eliminate-idempotent-choices r ⟩

eliminate-idempotent-choices-preserves
: {i : Size} {A : 𝔸}
(e : 2CC Dimension i A)
⟦ eliminate-idempotent-choices e ⟧ ≗ ⟦ e ⟧
eliminate-idempotent-choices-preserves (a -< cs >-) c =
⟦ eliminate-idempotent-choices (a -< cs >-) ⟧ c
≡⟨⟩
⟦ a -< List.map eliminate-idempotent-choices cs >- ⟧ c
≡⟨⟩
a V.-< List.map (λ e ⟦ e ⟧ c) (List.map eliminate-idempotent-choices cs) >-
≡⟨ Eq.cong (a Rose.-<_>-) (List.map-∘ cs) ⟨
a V.-< List.map (λ e ⟦ eliminate-idempotent-choices e ⟧ c) cs >-
≡⟨ Eq.cong (a Rose.-<_>-) (List.map-cong (λ e eliminate-idempotent-choices-preserves e c) cs) ⟩
a V.-< List.map (λ e ⟦ e ⟧ c) cs >-
≡⟨⟩
⟦ a -< cs >- ⟧ c
eliminate-idempotent-choices-preserves (D ⟨ l , r ⟩) c with eliminate-idempotent-choices l ≟ eliminate-idempotent-choices r
eliminate-idempotent-choices-preserves (D ⟨ l , r ⟩) c | no l≢r =
(if c D then ⟦ eliminate-idempotent-choices l ⟧ c else ⟦ eliminate-idempotent-choices r ⟧ c)
≡⟨ Eq.cong₂ (if c D then_else_) (eliminate-idempotent-choices-preserves l c) (eliminate-idempotent-choices-preserves r c) ⟩
(if c D then ⟦ l ⟧ c else ⟦ r ⟧ c)
≡⟨⟩
⟦ D ⟨ l , r ⟩ ⟧ c
eliminate-idempotent-choices-preserves (D ⟨ l , r ⟩) c | yes l≡r with c D
eliminate-idempotent-choices-preserves (D ⟨ l , r ⟩) c | yes l≡r | true =
⟦ eliminate-idempotent-choices l ⟧ c
≡⟨ eliminate-idempotent-choices-preserves l c ⟩
⟦ l ⟧ c
≡⟨⟩
(if true then ⟦ l ⟧ c else ⟦ r ⟧ c)
eliminate-idempotent-choices-preserves (D ⟨ l , r ⟩) c | yes l≡r | false =
⟦ eliminate-idempotent-choices l ⟧ c
≡⟨ Eq.cong₂ ⟦_⟧ l≡r refl ⟩
⟦ eliminate-idempotent-choices r ⟧ c
≡⟨ eliminate-idempotent-choices-preserves r c ⟩
⟦ r ⟧ c
≡⟨⟩
(if false then ⟦ l ⟧ c else ⟦ r ⟧ c)

Idempotence-Elimination : LanguageCompiler (2CCL Dimension) (2CCL Dimension)
Idempotence-Elimination = record
{ compile = eliminate-idempotent-choices
; config-compiler = λ _ record { to = id ; from = id }
; preserves = λ e ≅[]-sym (≗→≅[] (eliminate-idempotent-choices-preserves e))
}

0 comments on commit 81a5256

Please sign in to comment.