From 14b3318b16b089de4ee6eb3d7cbf6141bac33a7d Mon Sep 17 00:00:00 2001 From: pmbittner Date: Thu, 2 Nov 2023 20:52:17 +0100 Subject: [PATCH 01/43] first steps for feature algebra --- src/Framework/V2/Lang/FeatureAlgebra.agda | 85 +++++++++++++++++++++++ 1 file changed, 85 insertions(+) create mode 100644 src/Framework/V2/Lang/FeatureAlgebra.agda diff --git a/src/Framework/V2/Lang/FeatureAlgebra.agda b/src/Framework/V2/Lang/FeatureAlgebra.agda new file mode 100644 index 00000000..19a28207 --- /dev/null +++ b/src/Framework/V2/Lang/FeatureAlgebra.agda @@ -0,0 +1,85 @@ +module Framework.V2.Lang.FeatureAlgebra where + +open import Data.Product using (proj₁; proj₂; _,_) +open import Data.List using (List) renaming (_∷_ to _._) + +open import Algebra.Structures using (IsMonoid) +open import Algebra.Core using (Op₂) +import Algebra.Definitions + +open import Relation.Binary using (Rel; Reflexive; Transitive; IsEquivalence) +open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) + +open import Level using (0ℓ; suc; _⊔_) + +open import Framework.V2.Annotation.Name using (Name) + +record FeaturePath (N : Set) (S : Set) : Set where + constructor _∷_ + field + name : Name N + path : List S + +record FeatureAlgebra {c} : Set (suc c) where + open Algebra.Definitions using (Associative) + open Eq.≡-Reasoning + infixr 7 _⊕_ + + field + I : Set c + _⊕_ : Op₂ I + 𝟘 : I + + monoid : IsMonoid _≡_ _⊕_ 𝟘 + + -- Only the rightmost occurence of an introduction is effective in a sum, + -- because it has been introduced first. + -- This is, duplicates of i have no effect. + distant-idempotence : ∀ (i₁ i₂ : I) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ + direct-idempotence : ∀ (i : I) → i ⊕ i ≡ i + + open IsMonoid monoid + + -- introduction inclusion + infix 6 _≤_ + _≤_ : Rel I c + i₂ ≤ i₁ = i₂ ⊕ i₁ ≡ i₁ + + ≤-refl : Reflexive _≤_ + ≤-refl {i} = direct-idempotence i + + ≤-trans : Transitive _≤_ + ≤-trans {i} {j} {k} i≤j j≤k = + begin + i ⊕ k + ≡⟨ Eq.cong (i ⊕_) (Eq.sym j≤k) ⟩ + i ⊕ (j ⊕ k) + ≡⟨ Eq.cong (λ x → i ⊕ x ⊕ k) (Eq.sym i≤j) ⟩ + i ⊕ ((i ⊕ j) ⊕ k) + ≡⟨ Eq.sym (assoc i (i ⊕ j) k) ⟩ + (i ⊕ (i ⊕ j)) ⊕ k + ≡⟨ Eq.cong (_⊕ k) (Eq.sym (assoc i i j)) ⟩ + ((i ⊕ i) ⊕ j) ⊕ k + ≡⟨ Eq.cong (_⊕ k) (Eq.cong (_⊕ j) (direct-idempotence i)) ⟩ + (i ⊕ j) ⊕ k + ≡⟨ Eq.cong (_⊕ k) i≤j ⟩ + j ⊕ k + ≡⟨ j≤k ⟩ + k + ∎ + + least-element : ∀ i → 𝟘 ≤ i + least-element = proj₁ identity + + upper-bound-l : ∀ i₂ i₁ → i₂ ≤ i₂ ⊕ i₁ + upper-bound-l i₂ i₁ = + begin + i₂ ⊕ (i₂ ⊕ i₁) + ≡⟨ Eq.sym (assoc i₂ i₂ i₁) ⟩ + (i₂ ⊕ i₂) ⊕ i₁ + ≡⟨ Eq.cong (_⊕ i₁) (direct-idempotence i₂) ⟩ + i₂ ⊕ i₁ + ∎ + + upper-bound-r : ∀ i₂ i₁ → i₁ ≤ i₂ ⊕ i₁ + upper-bound-r i₂ i₁ = distant-idempotence i₂ i₁ From 133411918398c7644ddfbd630869104538c49cb3 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 11:27:09 +0100 Subject: [PATCH 02/43] direct-idempotence follows from distant-idemp --- src/Framework/V2/Lang/FeatureAlgebra.agda | 13 ++++++++++++- 1 file changed, 12 insertions(+), 1 deletion(-) diff --git a/src/Framework/V2/Lang/FeatureAlgebra.agda b/src/Framework/V2/Lang/FeatureAlgebra.agda index 19a28207..8f7092d7 100644 --- a/src/Framework/V2/Lang/FeatureAlgebra.agda +++ b/src/Framework/V2/Lang/FeatureAlgebra.agda @@ -36,10 +36,21 @@ record FeatureAlgebra {c} : Set (suc c) where -- because it has been introduced first. -- This is, duplicates of i have no effect. distant-idempotence : ∀ (i₁ i₂ : I) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ - direct-idempotence : ∀ (i : I) → i ⊕ i ≡ i open IsMonoid monoid + direct-idempotence : ∀ (i : I) → i ⊕ i ≡ i + direct-idempotence i = + begin + i ⊕ i + ≡˘⟨ Eq.cong (i ⊕_) (proj₁ identity i) ⟩ + i ⊕ 𝟘 ⊕ i + ≡⟨ distant-idempotence 𝟘 i ⟩ + 𝟘 ⊕ i + ≡⟨ proj₁ identity i ⟩ + i + ∎ + -- introduction inclusion infix 6 _≤_ _≤_ : Rel I c From 1ca124d335e087a123ae4a84580fae49ec076e56 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 11:27:46 +0100 Subject: [PATCH 03/43] sym refinements --- src/Framework/V2/Lang/FeatureAlgebra.agda | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/Framework/V2/Lang/FeatureAlgebra.agda b/src/Framework/V2/Lang/FeatureAlgebra.agda index 8f7092d7..19a64b89 100644 --- a/src/Framework/V2/Lang/FeatureAlgebra.agda +++ b/src/Framework/V2/Lang/FeatureAlgebra.agda @@ -63,13 +63,13 @@ record FeatureAlgebra {c} : Set (suc c) where ≤-trans {i} {j} {k} i≤j j≤k = begin i ⊕ k - ≡⟨ Eq.cong (i ⊕_) (Eq.sym j≤k) ⟩ + ≡˘⟨ Eq.cong (i ⊕_) j≤k ⟩ i ⊕ (j ⊕ k) - ≡⟨ Eq.cong (λ x → i ⊕ x ⊕ k) (Eq.sym i≤j) ⟩ + ≡˘⟨ Eq.cong (λ x → i ⊕ x ⊕ k) i≤j ⟩ i ⊕ ((i ⊕ j) ⊕ k) - ≡⟨ Eq.sym (assoc i (i ⊕ j) k) ⟩ + ≡˘⟨ assoc i (i ⊕ j) k ⟩ (i ⊕ (i ⊕ j)) ⊕ k - ≡⟨ Eq.cong (_⊕ k) (Eq.sym (assoc i i j)) ⟩ + ≡˘⟨ Eq.cong (_⊕ k) (assoc i i j) ⟩ ((i ⊕ i) ⊕ j) ⊕ k ≡⟨ Eq.cong (_⊕ k) (Eq.cong (_⊕ j) (direct-idempotence i)) ⟩ (i ⊕ j) ⊕ k From 1f427d496abffb9de3e7273eefeb95e9307ad652 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 11:44:14 +0100 Subject: [PATCH 04/43] finished FeatureAlgebra basic definition --- src/Framework/V2/Lang/FeatureAlgebra.agda | 74 ++++++++++++++++++++++- 1 file changed, 71 insertions(+), 3 deletions(-) diff --git a/src/Framework/V2/Lang/FeatureAlgebra.agda b/src/Framework/V2/Lang/FeatureAlgebra.agda index 19a64b89..f5f83be4 100644 --- a/src/Framework/V2/Lang/FeatureAlgebra.agda +++ b/src/Framework/V2/Lang/FeatureAlgebra.agda @@ -1,16 +1,16 @@ module Framework.V2.Lang.FeatureAlgebra where -open import Data.Product using (proj₁; proj₂; _,_) +open import Data.Product using (proj₁; proj₂; _×_; _,_) open import Data.List using (List) renaming (_∷_ to _._) open import Algebra.Structures using (IsMonoid) open import Algebra.Core using (Op₂) import Algebra.Definitions -open import Relation.Binary using (Rel; Reflexive; Transitive; IsEquivalence) +open import Relation.Binary using (Rel; Reflexive; Symmetric; Transitive; IsEquivalence) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) -open import Level using (0ℓ; suc; _⊔_) +open import Level using (suc; _⊔_) open import Framework.V2.Annotation.Name using (Name) @@ -79,9 +79,19 @@ record FeatureAlgebra {c} : Set (suc c) where k ∎ + ≤-IsPreorder : IsPreorder _≡_ _≤_ + ≤-IsPreorder = record + { isEquivalence = Eq.isEquivalence + ; reflexive = λ where refl → ≤-refl + ; trans = ≤-trans + } + least-element : ∀ i → 𝟘 ≤ i least-element = proj₁ identity + least-element-unique : ∀ i → i ≤ 𝟘 → i ≡ 𝟘 + least-element-unique i i≤𝟘 rewrite (proj₂ identity i) = i≤𝟘 + upper-bound-l : ∀ i₂ i₁ → i₂ ≤ i₂ ⊕ i₁ upper-bound-l i₂ i₁ = begin @@ -94,3 +104,61 @@ record FeatureAlgebra {c} : Set (suc c) where upper-bound-r : ∀ i₂ i₁ → i₁ ≤ i₂ ⊕ i₁ upper-bound-r i₂ i₁ = distant-idempotence i₂ i₁ + + least-upper-bound : ∀ i i₂ i₁ + → i₁ ≤ i + → i₂ ≤ i + ----------- + → i₁ ⊕ i₂ ≤ i + least-upper-bound i i₂ i₁ i₁≤i i₂≤i = + begin + (i₁ ⊕ i₂) ⊕ i + ≡⟨ assoc i₁ i₂ i ⟩ + i₁ ⊕ (i₂ ⊕ i) + ≡⟨ Eq.cong (i₁ ⊕_) i₂≤i ⟩ + i₁ ⊕ i + ≡⟨ i₁≤i ⟩ + i + ∎ + + -- introduction equivalence + infix 6 _~_ + _~_ : Rel I c + i₂ ~ i₁ = i₂ ≤ i₁ × i₁ ≤ i₂ + + ~-refl : Reflexive _~_ + ~-refl = ≤-refl , ≤-refl + + ~-sym : Symmetric _~_ + ~-sym (fst , snd) = snd , fst + + ~-trans : Transitive _~_ + ~-trans (i≤j , j≤i) (j≤k , k≤j) = ≤-trans i≤j j≤k , ≤-trans k≤j j≤i + + ~-IsEquivalence : IsEquivalence _~_ + ~-IsEquivalence = record + { refl = ~-refl + ; sym = ~-sym + ; trans = ~-trans + } + + quasi-smaller : ∀ i₂ i₁ → i₂ ⊕ i₁ ≤ i₁ ⊕ i₂ + quasi-smaller i₂ i₁ = + begin + (i₂ ⊕ i₁) ⊕ i₁ ⊕ i₂ + ≡⟨⟩ + (i₂ ⊕ i₁) ⊕ (i₁ ⊕ i₂) + ≡˘⟨ assoc (i₂ ⊕ i₁) i₁ i₂ ⟩ + ((i₂ ⊕ i₁) ⊕ i₁) ⊕ i₂ + ≡⟨ Eq.cong (_⊕ i₂) (assoc i₂ i₁ i₁) ⟩ + (i₂ ⊕ (i₁ ⊕ i₁)) ⊕ i₂ + ≡⟨ Eq.cong (_⊕ i₂) (Eq.cong (i₂ ⊕_) (direct-idempotence i₁)) ⟩ + (i₂ ⊕ i₁) ⊕ i₂ + ≡⟨ assoc i₂ i₁ i₂ ⟩ + i₂ ⊕ i₁ ⊕ i₂ + ≡⟨ distant-idempotence i₁ i₂ ⟩ + i₁ ⊕ i₂ + ∎ + + quasi-commutativity : ∀ i₂ i₁ → i₂ ⊕ i₁ ~ i₁ ⊕ i₂ + quasi-commutativity i₂ i₁ = quasi-smaller i₂ i₁ , quasi-smaller i₁ i₂ From c93e690e902179764e62f547c0197bba28f3fd18 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 12:05:26 +0100 Subject: [PATCH 05/43] refactor: Extract FSTs --- src/Framework/V2/Lang/FST.agda | 11 +++++++++++ src/Framework/V2/Lang/FeatureAlgebra.agda | 17 ++--------------- 2 files changed, 13 insertions(+), 15 deletions(-) create mode 100644 src/Framework/V2/Lang/FST.agda diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda new file mode 100644 index 00000000..02788dfd --- /dev/null +++ b/src/Framework/V2/Lang/FST.agda @@ -0,0 +1,11 @@ +module Framework.V2.Lang.FST where + +open import Data.List using (List) renaming (_∷_ to _._) + +open import Framework.V2.Annotation.Name using (Name) + +record FeaturePath (N : Set) (S : Set) : Set where + constructor _∷_ + field + name : Name N + path : List S diff --git a/src/Framework/V2/Lang/FeatureAlgebra.agda b/src/Framework/V2/Lang/FeatureAlgebra.agda index f5f83be4..92b58b64 100644 --- a/src/Framework/V2/Lang/FeatureAlgebra.agda +++ b/src/Framework/V2/Lang/FeatureAlgebra.agda @@ -1,27 +1,14 @@ module Framework.V2.Lang.FeatureAlgebra where open import Data.Product using (proj₁; proj₂; _×_; _,_) -open import Data.List using (List) renaming (_∷_ to _._) - open import Algebra.Structures using (IsMonoid) open import Algebra.Core using (Op₂) -import Algebra.Definitions - -open import Relation.Binary using (Rel; Reflexive; Symmetric; Transitive; IsEquivalence) +open import Algebra.Definitions using (Associative) +open import Relation.Binary using (Rel; Reflexive; Symmetric; Transitive; IsEquivalence; IsPreorder) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) - open import Level using (suc; _⊔_) -open import Framework.V2.Annotation.Name using (Name) - -record FeaturePath (N : Set) (S : Set) : Set where - constructor _∷_ - field - name : Name N - path : List S - record FeatureAlgebra {c} : Set (suc c) where - open Algebra.Definitions using (Associative) open Eq.≡-Reasoning infixr 7 _⊕_ From 0607103baab97229e810b33fcc67ba776f3da3f4 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 13:05:41 +0100 Subject: [PATCH 06/43] refinements in FeatureAlgebra --- src/Framework/V2/Lang/FeatureAlgebra.agda | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/src/Framework/V2/Lang/FeatureAlgebra.agda b/src/Framework/V2/Lang/FeatureAlgebra.agda index 92b58b64..12911b05 100644 --- a/src/Framework/V2/Lang/FeatureAlgebra.agda +++ b/src/Framework/V2/Lang/FeatureAlgebra.agda @@ -8,15 +8,13 @@ open import Relation.Binary using (Rel; Reflexive; Symmetric; Transitive; IsEqui open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) open import Level using (suc; _⊔_) -record FeatureAlgebra {c} : Set (suc c) where +record FeatureAlgebra {c} (I : Set c) (sum : Op₂ I) (𝟘 : I) : Set (suc c) where open Eq.≡-Reasoning + + _⊕_ = sum infixr 7 _⊕_ field - I : Set c - _⊕_ : Op₂ I - 𝟘 : I - monoid : IsMonoid _≡_ _⊕_ 𝟘 -- Only the rightmost occurence of an introduction is effective in a sum, From ab2bd1ce1599af005f9d9f6387dd7abc8b52f5ce Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 13:05:50 +0100 Subject: [PATCH 07/43] first definition of FSTs --- src/Framework/V2/Lang/FST.agda | 62 ++++++++++++++++++++++++++++++---- 1 file changed, 55 insertions(+), 7 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 02788dfd..a56b531f 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -1,11 +1,59 @@ -module Framework.V2.Lang.FST where +open import Framework.V2.Definitions +open import Relation.Binary using (DecidableEquality) -open import Data.List using (List) renaming (_∷_ to _._) +module Framework.V2.Lang.FST + (A : 𝔸) + (_≟_ : DecidableEquality A) + where +open import Data.Bool using (Bool; true; false) +open import Data.List using (List; []; _∷_; foldr; map) + +open import Relation.Nullary.Decidable using (_because_) + +open import Framework.V2.Variants using (Rose; artifact) open import Framework.V2.Annotation.Name using (Name) +open import Framework.V2.Constructs.Artifact +open import Framework.V2.Lang.FeatureAlgebra + +FeaturePath : Set → Set +FeaturePath = List + +data FST : 𝕍 where + root : ∀ {A : 𝔸} → List (Rose A) → FST A + +FeatureForest : (N : Set) → 𝔼 +FeatureForest N A = ∀ (n : Name N) → FeaturePath A + +𝟘 : FST A +𝟘 = root [] + +-- We could avoid wrap and unwrap by defining our own intermediate tree structure +-- that does not reuse Artifact constructor. +unwrap : Rose A → Artifact A (Rose A) +unwrap (artifact a) = a + +wrap : Artifact A (Rose A) → Rose A +wrap a = artifact a + +mutual + -- TODO: Avoid termination macro. + {-# TERMINATING #-} + impose-subtree : Artifact A (Rose A) → List (Artifact A (Rose A)) → List (Artifact A (Rose A)) + impose-subtree l [] = l ∷ [] + impose-subtree (a -< as >-) (b -< bs >- ∷ rs) with a ≟ b + ... | true because _ = b -< impose as bs >- ∷ rs + ... | false because _ = impose-subtree (a -< as >-) rs + + impose-raw : List (Artifact A (Rose A)) → List (Artifact A (Rose A)) → List (Artifact A (Rose A)) + impose-raw ls rs = foldr impose-subtree rs ls + + impose : List (Rose A) → List (Rose A) → List (Rose A) + impose ls rs = map wrap (impose-raw (map unwrap ls) (map unwrap rs)) + +infixr 7 _⊕_ +_⊕_ : FST A → FST A → FST A +root l ⊕ root r = root (impose l r) -record FeaturePath (N : Set) (S : Set) : Set where - constructor _∷_ - field - name : Name N - path : List S +-- FST-is-FeatureAlgebra : FeatureAlgebra (FST A) _⊕_ 𝟘 +-- FST-is-FeatureAlgebra = {!!} From 578ce4cb8636a4dc124c8ace9e44ce4d34ce9fc4 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 18:05:26 +0100 Subject: [PATCH 08/43] FST refinements --- src/Framework/V2/Lang/FST.agda | 128 +++++++++++++++++++++++---------- 1 file changed, 89 insertions(+), 39 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index a56b531f..276a38d0 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -1,59 +1,109 @@ -open import Framework.V2.Definitions -open import Relation.Binary using (DecidableEquality) - -module Framework.V2.Lang.FST - (A : 𝔸) - (_≟_ : DecidableEquality A) - where +module Framework.V2.Lang.FST where -open import Data.Bool using (Bool; true; false) -open import Data.List using (List; []; _∷_; foldr; map) +open import Data.Bool using (Bool; true; false; if_then_else_) +open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) +open import Function using (_∘_) -open import Relation.Nullary.Decidable using (_because_) +open import Relation.Binary using (DecidableEquality) +open import Relation.Nullary.Decidable using (yes; no) -open import Framework.V2.Variants using (Rose; artifact) +open import Framework.V2.Definitions +open import Framework.V2.Variants using (artifact) open import Framework.V2.Annotation.Name using (Name) open import Framework.V2.Constructs.Artifact open import Framework.V2.Lang.FeatureAlgebra -FeaturePath : Set → Set -FeaturePath = List +data FSTNode : 𝕍 where + node : ∀ {A} → A → List (FSTNode A) → FSTNode A + +-- All FSTs have the same implicit root. +-- So an FST is just a list of children, implicitly grouped below +-- an imaginary unique root. +FST : 𝕍 +FST A = List (FSTNode A) + +infixr 4 _::_ +record Feature (N : 𝔽) (A : 𝔸) : Set where + constructor _::_ + field + name : Name N + impl : FST A +open Feature public + +-- the syntax used in the paper for paths +infixr 5 _._ +_._ : ∀ {A : 𝔸} → A → List (FSTNode A) → FST A +a . cs = node a cs ∷ [] + +-- helper function when branching in paths +branches : ∀ {A : 𝔸} → List (List (FSTNode A)) → List (FSTNode A) +branches = concat -data FST : 𝕍 where - root : ∀ {A : 𝔸} → List (Rose A) → FST A +FeatureForest : (N : 𝔽) → 𝔼 +FeatureForest N A = List (Feature N A) -FeatureForest : (N : Set) → 𝔼 -FeatureForest N A = ∀ (n : Name N) → FeaturePath A +Conf : (N : 𝔽) → Set +Conf N = Config N Bool -𝟘 : FST A -𝟘 = root [] +select : ∀ {N A} → Conf N → FeatureForest N A → FeatureForest N A +select c = filterᵇ (c ∘ name) --- We could avoid wrap and unwrap by defining our own intermediate tree structure --- that does not reuse Artifact constructor. -unwrap : Rose A → Artifact A (Rose A) -unwrap (artifact a) = a +forget-names : ∀ {N A} → FeatureForest N A → List (FST A) +forget-names = map impl -wrap : Artifact A (Rose A) → Rose A -wrap a = artifact a +names : ∀ {N A} → FeatureForest N A → List N +names = map name -mutual +module Algebra {A : 𝔸} (_≟_ : DecidableEquality A) where + 𝟘 : FST A + 𝟘 = [] + + mutual -- TODO: Avoid termination macro. {-# TERMINATING #-} - impose-subtree : Artifact A (Rose A) → List (Artifact A (Rose A)) → List (Artifact A (Rose A)) + impose-subtree : FSTNode A → List (FSTNode A) → List (FSTNode A) impose-subtree l [] = l ∷ [] - impose-subtree (a -< as >-) (b -< bs >- ∷ rs) with a ≟ b - ... | true because _ = b -< impose as bs >- ∷ rs - ... | false because _ = impose-subtree (a -< as >-) rs + impose-subtree (node a as) (node b bs ∷ rs) with a ≟ b + ... | yes _ = node b (as ⊕ bs) ∷ rs + ... | no _ = node b bs ∷ impose-subtree (node a as) rs + + infixr 7 _⊕_ + _⊕_ : FST A → FST A → FST A + l ⊕ r = foldr impose-subtree r l + + ⊕-all : List (FST A) → FST A + ⊕-all = foldr _⊕_ 𝟘 + + -- FST-is-FeatureAlgebra : FeatureAlgebra (FST A) _⊕_ 𝟘 + -- FST-is-FeatureAlgebra = {!!} - impose-raw : List (Artifact A (Rose A)) → List (Artifact A (Rose A)) → List (Artifact A (Rose A)) - impose-raw ls rs = foldr impose-subtree rs ls + ⟦_⟧ : ∀ {N : 𝔽} → FeatureForest N A → Conf N → FST A + ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features - impose : List (Rose A) → List (Rose A) → List (Rose A) - impose ls rs = map wrap (impose-raw (map unwrap ls) (map unwrap rs)) + -- We could avoid wrap and unwrap by defining our own intermediate tree structure + -- that does not reuse Artifact constructor. + -- unwrap : Rose A → Artifact A (Rose A) + -- unwrap (artifact a) = a + + -- wrap : Artifact A (Rose A) → Rose A + -- wrap a = artifact a + +open import Data.String using (String; _<+>_) +open import Show.Lines + +module Show {N : 𝔽} {A : 𝔸} (show-N : N → String) (show-A : A → String) where + mutual + -- TODO: Why does termination checking fail here? + {-# TERMINATING #-} + show-FSTNode : FSTNode A → Lines + show-FSTNode (node a children) = do + > show-A a + indent 2 (show-FST children) -infixr 7 _⊕_ -_⊕_ : FST A → FST A → FST A -root l ⊕ root r = root (impose l r) + show-FST : FST A → Lines + show-FST fst = lines (map show-FSTNode fst) --- FST-is-FeatureAlgebra : FeatureAlgebra (FST A) _⊕_ 𝟘 --- FST-is-FeatureAlgebra = {!!} + show-Feature : Feature N A → Lines + show-Feature feature = do + > show-N (name feature) <+> "∷" + indent 2 (show-FST (impl feature)) From 1e302e8ca6dbabff07793220b304198798f97746 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 18:05:36 +0100 Subject: [PATCH 09/43] map-named --- src/Util/Named.agda | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/Util/Named.agda b/src/Util/Named.agda index 0c4267e6..cc3f4aa2 100644 --- a/src/Util/Named.agda +++ b/src/Util/Named.agda @@ -13,5 +13,8 @@ record Named {ℓ} (A : Set ℓ) : Set ℓ where open Named public infix 2 _called_ +map-named : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} → (A → B) → Named A → Named B +map-named f (get called name) = f get called name + show-named : ∀ {ℓ} {A : Set ℓ} → (A → String) → Named A → String show-named show-A n = (getName n) ++ " = " ++ show-A (get n) From 6d46808b1c27615684938e1b1c11dbb1e58ab082 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 18:05:43 +0100 Subject: [PATCH 10/43] for-loop in pretty printing --- src/Show/Lines.agda | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/Show/Lines.agda b/src/Show/Lines.agda index d164b1a5..087a092b 100644 --- a/src/Show/Lines.agda +++ b/src/Show/Lines.agda @@ -47,6 +47,11 @@ single = pure list-applicative lines : List Lines → Lines lines = concat +for-loop : ∀ {ℓ} {A : Set ℓ} → List A → (A → Lines) → Lines +for-loop items op = lines (map op items) + +syntax for-loop items (λ c → l) = foreach [ c ∈ items ] l + align-all : ℕ → Lines → Lines align-all width = map (align width) From 4f84cb6007f16d807e3dbc25f7dcaacdd926a376 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 18:05:59 +0100 Subject: [PATCH 11/43] experiment for FSTs --- .../V2/Experiment/FST-Experiments.agda | 111 ++++++++++++++++++ src/Main.agda | 3 + 2 files changed, 114 insertions(+) create mode 100644 src/Framework/V2/Experiment/FST-Experiments.agda diff --git a/src/Framework/V2/Experiment/FST-Experiments.agda b/src/Framework/V2/Experiment/FST-Experiments.agda new file mode 100644 index 00000000..d062eee3 --- /dev/null +++ b/src/Framework/V2/Experiment/FST-Experiments.agda @@ -0,0 +1,111 @@ +open import Framework.V2.Definitions +module Framework.V2.Experiment.FST-Experiments where + +open import Data.Bool using (true; false) +open import Data.List using (List; _∷_; []; map; [_]) +open import Data.Product using (proj₁; proj₂; _,_; _×_) +open import Function using (id) + +open import Relation.Binary using (DecidableEquality) +open import Relation.Nullary.Decidable using (yes; no; does; _because_; _×-dec_) + +open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) + +open import Framework.V2.Lang.FST as FSTModule + +open import Util.Named +open import Test.Example +open import Test.Experiment +open import Show.Lines +open import Util.ShowHelpers +open import Data.String using (String; _<+>_; _++_; _≟_) + +exp : ∀ {N A} + → (N → String) + → (A → String) + → DecidableEquality A + → List (Conf N) + → Experiment (FeatureForest N A) +getName (exp _ _ _ _) = "Configure FST example" +get (exp show-N show-A _≟_ configs) (example-name ≔ forest) = + let open FSTModule.Show show-N show-A + open FSTModule.Algebra _≟_ + in + do + > "Expression e has features" + indent 2 do + lines (map show-Feature forest) + + foreach [ c ∈ configs ] do + let cstr = show-fun show-N show-bool c (names forest) + linebreak + > "⟦ e ⟧" <+> cstr <+> "=" + indent 2 do + show-FST (⟦ forest ⟧ c) + +pick-all : ∀ {N} → Conf N +pick-all _ = true + +pick-only : ∀ {N} → DecidableEquality N → N → Conf N +pick-only _==_ n n' = does (n == n') + +module Java where + ASTNode = String + + class = "class" <+>_ + method = "method" <+>_ + return : String → ASTNode + return e = "return" <+> e ++ ";" + + _≟-ast_ : DecidableEquality ASTNode + _≟-ast_ = _≟_ + + module Calculator where + fname-Add = "Add" + fname-Sub = "Sub" + fname-Log = "Log" + + cls = class "Calculator" + + add = method "add(int,int)" + add-ret = return "x + y" + + sub = method "sub(int,int)" + sub-ret = return "x - y" + + log = "log(x + \", \" + y);" + + ---- Features + + feature-Add : Feature String ASTNode + feature-Add = fname-Add :: cls . add . add-ret . [] + + feature-Sub : Feature String ASTNode + feature-Sub = fname-Sub :: cls . sub . sub-ret . [] + + feature-Log : Feature String ASTNode + feature-Log = fname-Log :: cls . + branches ( + (add . log . []) + ∷ (sub . log . []) + ∷ []) + + ---- Example SPLs + + ex-Add-Sub : Example (FeatureForest String ASTNode) + ex-Add-Sub = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ [] + + ex-Sub-Add : Example (FeatureForest String ASTNode) + ex-Sub-Add = "sub-add" ≔ feature-Sub ∷ feature-Add ∷ [] + + ex-Add-Sub-Log : Example (FeatureForest String ASTNode) + ex-Add-Sub-Log = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ feature-Log ∷ [] + + ex-all : List (Example (FeatureForest String ASTNode)) + ex-all = ex-Add-Sub ∷ ex-Sub-Add ∷ ex-Add-Sub-Log ∷ [] + + ---- Experiments + + toy-calculator-experiment = + let eq = _≟-ast_ in + exp id id eq (pick-all ∷ pick-only eq fname-Add ∷ []) diff --git a/src/Main.agda b/src/Main.agda index 107fd2f5..f9e7c3db 100644 --- a/src/Main.agda +++ b/src/Main.agda @@ -25,6 +25,8 @@ open import Test.Experiments.CCC-to-BCC open import Test.Experiments.OC-to-BCC open import Framework.V2.Translation.Experiments.NChoice-to-2Choice-Experiment using (exp; all-ex) +import Framework.V2.Experiment.FST-Experiments as FSTs +open FSTs.Java.Calculator using (toy-calculator-experiment; ex-all) {-| A list of programs that we want to run. @@ -39,6 +41,7 @@ experimentsToRun = setup exp-oc-to-bcc optex-all ∷ -- Run some example translations from b-ary to binary choices setup exp all-ex ∷ + setup toy-calculator-experiment ex-all ∷ [] {-| From 0b7383a8f6475be33e28e13ddcf1e55f5ce4cf6f Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 18:33:19 +0100 Subject: [PATCH 12/43] progress and observation --- src/Framework/V2/Lang/FST.agda | 85 +++++++++++++++++++++++++++------- 1 file changed, 68 insertions(+), 17 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 276a38d0..965721c6 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -4,8 +4,9 @@ open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) open import Function using (_∘_) -open import Relation.Binary using (DecidableEquality) open import Relation.Nullary.Decidable using (yes; no) +open import Relation.Binary using (DecidableEquality) +open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) open import Framework.V2.Definitions open import Framework.V2.Variants using (artifact) @@ -55,6 +56,9 @@ names : ∀ {N A} → FeatureForest N A → List N names = map name module Algebra {A : 𝔸} (_≟_ : DecidableEquality A) where + open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) + open Eq.≡-Reasoning + 𝟘 : FST A 𝟘 = [] @@ -71,22 +75,69 @@ module Algebra {A : 𝔸} (_≟_ : DecidableEquality A) where _⊕_ : FST A → FST A → FST A l ⊕ r = foldr impose-subtree r l - ⊕-all : List (FST A) → FST A - ⊕-all = foldr _⊕_ 𝟘 - - -- FST-is-FeatureAlgebra : FeatureAlgebra (FST A) _⊕_ 𝟘 - -- FST-is-FeatureAlgebra = {!!} - - ⟦_⟧ : ∀ {N : 𝔽} → FeatureForest N A → Conf N → FST A - ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features - - -- We could avoid wrap and unwrap by defining our own intermediate tree structure - -- that does not reuse Artifact constructor. - -- unwrap : Rose A → Artifact A (Rose A) - -- unwrap (artifact a) = a - - -- wrap : Artifact A (Rose A) → Rose A - -- wrap a = artifact a + ⊕-all : List (FST A) → FST A + ⊕-all = foldr _⊕_ 𝟘 + + l-id : LeftIdentity _≡_ 𝟘 _⊕_ + l-id _ = refl + + -- This is not satisfied. What did we do wrong? + -- I think the problem is that (x ∷ xs) ⊕ 𝟘 + -- denotes an FST superimposition of x onto xs, recursively, + -- which is not what we want. + -- What happens is that + -- 1.) x gets imposed onto 𝟘 and yields x + -- 2.) the next child in xs gets imposed onto x, potentially mutating x. + -- BOOM + -- TODO: How to fix that? This self-imposition also occurs when the rhs is not 𝟘. + -- So it is normal, right? + -- Maybe, the imposition should not be done sequentially but in parallel? + r-id : RightIdentity _≡_ 𝟘 _⊕_ + r-id [] = refl + r-id (x ∷ xs) = {!!} + -- rewrite r-id xs = + -- begin + -- impose-subtree x xs + -- ≡⟨ {!!} ⟩ + -- x ∷ xs + -- ∎ + + assoc : Associative _≡_ _⊕_ + assoc x y z = {!!} + + cong : Congruent₂ _≡_ _⊕_ + cong refl refl = refl + + idem : ∀ (i₁ i₂ : FST A) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ + idem = {!!} + + FST-is-FeatureAlgebra : FeatureAlgebra (FST A) _⊕_ 𝟘 + FST-is-FeatureAlgebra = record + { monoid = record + { isSemigroup = record + { isMagma = record + { isEquivalence = Eq.isEquivalence + ; ∙-cong = cong + } + ; assoc = assoc + } + ; identity = l-id , r-id + } + ; distant-idempotence = idem + } + where + open import Data.Product using (_,_) + + ⟦_⟧ : ∀ {N : 𝔽} → FeatureForest N A → Conf N → FST A + ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features + + -- We could avoid wrap and unwrap by defining our own intermediate tree structure + -- that does not reuse Artifact constructor. + -- unwrap : Rose A → Artifact A (Rose A) + -- unwrap (artifact a) = a + + -- wrap : Artifact A (Rose A) → Rose A + -- wrap a = artifact a open import Data.String using (String; _<+>_) open import Show.Lines From 73a30b038dc005bcde756bb838012623b74ccef2 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 3 Nov 2023 18:35:55 +0100 Subject: [PATCH 13/43] --allow-unsolved-metas --- src/Framework/V2/Experiment/FST-Experiments.agda | 2 ++ src/Framework/V2/Lang/FST.agda | 2 ++ src/Main.agda | 2 +- 3 files changed, 5 insertions(+), 1 deletion(-) diff --git a/src/Framework/V2/Experiment/FST-Experiments.agda b/src/Framework/V2/Experiment/FST-Experiments.agda index d062eee3..dafc4835 100644 --- a/src/Framework/V2/Experiment/FST-Experiments.agda +++ b/src/Framework/V2/Experiment/FST-Experiments.agda @@ -1,3 +1,5 @@ +{-# OPTIONS --allow-unsolved-metas #-} + open import Framework.V2.Definitions module Framework.V2.Experiment.FST-Experiments where diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 965721c6..53701638 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -1,3 +1,5 @@ +{-# OPTIONS --allow-unsolved-metas #-} + module Framework.V2.Lang.FST where open import Data.Bool using (Bool; true; false; if_then_else_) diff --git a/src/Main.agda b/src/Main.agda index f9e7c3db..594ecdfe 100644 --- a/src/Main.agda +++ b/src/Main.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --sized-types --guardedness #-} +{-# OPTIONS --sized-types --guardedness --allow-unsolved-metas #-} module Main where From a64b933997cb627ae45b8b4fe457eada04b817f1 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Mon, 6 Nov 2023 14:26:18 +0100 Subject: [PATCH 14/43] FST node children are disjunct artifacts --- .../V2/Experiment/FST-Experiments.agda | 87 ++++++----- src/Framework/V2/Lang/FST.agda | 138 ++++++++++-------- 2 files changed, 123 insertions(+), 102 deletions(-) diff --git a/src/Framework/V2/Experiment/FST-Experiments.agda b/src/Framework/V2/Experiment/FST-Experiments.agda index dafc4835..2e5ad3a8 100644 --- a/src/Framework/V2/Experiment/FST-Experiments.agda +++ b/src/Framework/V2/Experiment/FST-Experiments.agda @@ -13,37 +13,41 @@ open import Relation.Nullary.Decidable using (yes; no; does; _because_; _×-dec_ open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) -open import Framework.V2.Lang.FST as FSTModule - open import Util.Named open import Test.Example open import Test.Experiment open import Show.Lines open import Util.ShowHelpers -open import Data.String using (String; _<+>_; _++_; _≟_) - -exp : ∀ {N A} - → (N → String) - → (A → String) - → DecidableEquality A - → List (Conf N) - → Experiment (FeatureForest N A) -getName (exp _ _ _ _) = "Configure FST example" -get (exp show-N show-A _≟_ configs) (example-name ≔ forest) = - let open FSTModule.Show show-N show-A - open FSTModule.Algebra _≟_ - in - do - > "Expression e has features" - indent 2 do - lines (map show-Feature forest) - - foreach [ c ∈ configs ] do - let cstr = show-fun show-N show-bool c (names forest) - linebreak - > "⟦ e ⟧" <+> cstr <+> "=" +open import Data.String using (String; _<+>_; _++_) renaming (_≟_ to _≟ˢ_) + +open import Framework.V2.Lang.FST as FSTModule +module FSTDefs = FSTModule.Defs + +module _ {A : 𝔸} (_≟_ : DecidableEquality A) where + module FSTDefsA = FSTDefs _≟_ + open FSTDefsA + open FSTModule using (Conf) + + exp : ∀ {N} + → (N → String) + → (A → String) + → List (Conf N) + → Experiment (FSF N) + getName (exp _ _ _) = "Configure FST example" + get (exp show-N show-A configs) (example-name ≔ forest) = + let open FSTDefsA.Show show-N show-A + in + do + > "Expression e has features" indent 2 do - show-FST (⟦ forest ⟧ c) + lines (map show-Feature forest) + + foreach [ c ∈ configs ] do + let cstr = show-fun show-N show-bool c (names forest) + linebreak + > "⟦ e ⟧" <+> cstr <+> "=" + indent 2 do + show-FST (⟦ forest ⟧ c) pick-all : ∀ {N} → Conf N pick-all _ = true @@ -60,7 +64,9 @@ module Java where return e = "return" <+> e ++ ";" _≟-ast_ : DecidableEquality ASTNode - _≟-ast_ = _≟_ + _≟-ast_ = _≟ˢ_ + + open FSTDefs _≟-ast_ module Calculator where fname-Add = "Add" @@ -79,35 +85,38 @@ module Java where ---- Features - feature-Add : Feature String ASTNode - feature-Add = fname-Add :: cls . add . add-ret . [] + open import Data.Unit using (tt) + open import Data.List.Relation.Unary.AllPairs using ([]; _∷_) + open import Data.List.Relation.Unary.All using ([]; _∷_) + feature-Add : Feature ASTNode + feature-Add = fname-Add :: cls . add . add-ret . [] [ [] ] [ [] ∷ [] ] [ [] ∷ [] ] - feature-Sub : Feature String ASTNode - feature-Sub = fname-Sub :: cls . sub . sub-ret . [] + feature-Sub : Feature ASTNode + feature-Sub = fname-Sub :: cls . sub . sub-ret . [] [ [] ] [ [] ∷ [] ] [ [] ∷ [] ] - feature-Log : Feature String ASTNode + feature-Log : Feature ASTNode feature-Log = fname-Log :: cls . branches ( - (add . log . []) - ∷ (sub . log . []) - ∷ []) + (add . log . [] [ [] ] [ [] ∷ [] ]) + ∷ (sub . log . [] [ [] ] [ [] ∷ [] ]) + ∷ []) [ (tt ∷ []) ∷ [] ∷ [] ] ---- Example SPLs - ex-Add-Sub : Example (FeatureForest String ASTNode) + ex-Add-Sub : Example (FSF ASTNode) ex-Add-Sub = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ [] - ex-Sub-Add : Example (FeatureForest String ASTNode) + ex-Sub-Add : Example (FSF ASTNode) ex-Sub-Add = "sub-add" ≔ feature-Sub ∷ feature-Add ∷ [] - ex-Add-Sub-Log : Example (FeatureForest String ASTNode) + ex-Add-Sub-Log : Example (FSF ASTNode) ex-Add-Sub-Log = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ feature-Log ∷ [] - ex-all : List (Example (FeatureForest String ASTNode)) + ex-all : List (Example (FSF ASTNode)) ex-all = ex-Add-Sub ∷ ex-Sub-Add ∷ ex-Add-Sub-Log ∷ [] ---- Experiments toy-calculator-experiment = let eq = _≟-ast_ in - exp id id eq (pick-all ∷ pick-only eq fname-Add ∷ []) + exp eq id id (pick-all ∷ pick-only eq fname-Add ∷ []) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 53701638..798d90b1 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -4,10 +4,12 @@ module Framework.V2.Lang.FST where open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) +open import Data.List.Relation.Unary.AllPairs using (AllPairs) open import Function using (_∘_) +open import Level using (0ℓ) -open import Relation.Nullary.Decidable using (yes; no) -open import Relation.Binary using (DecidableEquality) +open import Relation.Nullary.Decidable using (yes; no; False) +open import Relation.Binary using (DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) open import Framework.V2.Definitions @@ -16,68 +18,78 @@ open import Framework.V2.Annotation.Name using (Name) open import Framework.V2.Constructs.Artifact open import Framework.V2.Lang.FeatureAlgebra -data FSTNode : 𝕍 where - node : ∀ {A} → A → List (FSTNode A) → FSTNode A +-- data FSTNode : 𝕍 where +-- -- add a proof that the children are disjoint wrt. ≡ +-- node : ∀ {A} → A → List (FSTNode A) → FSTNode A --- All FSTs have the same implicit root. --- So an FST is just a list of children, implicitly grouped below --- an imaginary unique root. -FST : 𝕍 -FST A = List (FSTNode A) +Conf : (N : 𝔽) → Set +Conf N = Config N Bool -infixr 4 _::_ -record Feature (N : 𝔽) (A : 𝔸) : Set where - constructor _::_ - field - name : Name N - impl : FST A -open Feature public +module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where + data FSTNode : Set + different : Rel FSTNode 0ℓ --- the syntax used in the paper for paths -infixr 5 _._ -_._ : ∀ {A : 𝔸} → A → List (FSTNode A) → FST A -a . cs = node a cs ∷ [] + data FSTNode where + node : A → (children : List FSTNode) → AllPairs different children → FSTNode --- helper function when branching in paths -branches : ∀ {A : 𝔸} → List (List (FSTNode A)) → List (FSTNode A) -branches = concat + different (node a _ _) (node b _ _) = False (a ≟ b) -FeatureForest : (N : 𝔽) → 𝔼 -FeatureForest N A = List (Feature N A) + -- All FSTs have the same implicit root. + -- So an FST is just a list of children, implicitly grouped below + -- an imaginary unique root. + FST : Set + FST = List FSTNode -Conf : (N : 𝔽) → Set -Conf N = Config N Bool + infixr 4 _::_ + record Feature (N : 𝔽) : Set where + constructor _::_ + field + name : Name N + impl : FST + open Feature public + +-- the syntax used in the paper for paths + infixr 5 _._[_] + _._[_] : A → (cs : List FSTNode) → AllPairs different cs → FST + a . cs [ d ] = node a cs d ∷ [] + + -- helper function when branching in paths + branches : List (List FSTNode) → List FSTNode + branches = concat + + -- Feature Structure Forest + FSF : (N : 𝔽) → Set --𝔼 + FSF N = List (Feature N) -select : ∀ {N A} → Conf N → FeatureForest N A → FeatureForest N A -select c = filterᵇ (c ∘ name) + select : ∀ {N} → Conf N → FSF N → FSF N + select c = filterᵇ (c ∘ name) -forget-names : ∀ {N A} → FeatureForest N A → List (FST A) -forget-names = map impl + forget-names : ∀ {N} → FSF N → List FST + forget-names = map impl -names : ∀ {N A} → FeatureForest N A → List N -names = map name + names : ∀ {N} → FSF N → List N + names = map name -module Algebra {A : 𝔸} (_≟_ : DecidableEquality A) where open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) open Eq.≡-Reasoning - 𝟘 : FST A + 𝟘 : FST 𝟘 = [] mutual -- TODO: Avoid termination macro. {-# TERMINATING #-} - impose-subtree : FSTNode A → List (FSTNode A) → List (FSTNode A) + impose-subtree : FSTNode → List FSTNode → List FSTNode impose-subtree l [] = l ∷ [] - impose-subtree (node a as) (node b bs ∷ rs) with a ≟ b - ... | yes _ = node b (as ⊕ bs) ∷ rs - ... | no _ = node b bs ∷ impose-subtree (node a as) rs + impose-subtree (node a as as-unique) (node b bs bs-unique ∷ rs) with a ≟ b + ... | yes _ = node b (as ⊕ bs) {!!} ∷ rs + ... | no _ = node b bs bs-unique ∷ impose-subtree (node a as as-unique) rs infixr 7 _⊕_ - _⊕_ : FST A → FST A → FST A + _⊕_ : FST → FST → FST l ⊕ r = foldr impose-subtree r l - ⊕-all : List (FST A) → FST A + ⊕-all : List FST → FST ⊕-all = foldr _⊕_ 𝟘 l-id : LeftIdentity _≡_ 𝟘 _⊕_ @@ -110,10 +122,10 @@ module Algebra {A : 𝔸} (_≟_ : DecidableEquality A) where cong : Congruent₂ _≡_ _⊕_ cong refl refl = refl - idem : ∀ (i₁ i₂ : FST A) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ + idem : ∀ (i₁ i₂ : FST) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ idem = {!!} - FST-is-FeatureAlgebra : FeatureAlgebra (FST A) _⊕_ 𝟘 + FST-is-FeatureAlgebra : FeatureAlgebra FST _⊕_ 𝟘 FST-is-FeatureAlgebra = record { monoid = record { isSemigroup = record @@ -130,7 +142,7 @@ module Algebra {A : 𝔸} (_≟_ : DecidableEquality A) where where open import Data.Product using (_,_) - ⟦_⟧ : ∀ {N : 𝔽} → FeatureForest N A → Conf N → FST A + ⟦_⟧ : ∀ {N : 𝔽} → FSF N → Conf N → FST ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features -- We could avoid wrap and unwrap by defining our own intermediate tree structure @@ -141,22 +153,22 @@ module Algebra {A : 𝔸} (_≟_ : DecidableEquality A) where -- wrap : Artifact A (Rose A) → Rose A -- wrap a = artifact a -open import Data.String using (String; _<+>_) -open import Show.Lines - -module Show {N : 𝔽} {A : 𝔸} (show-N : N → String) (show-A : A → String) where - mutual - -- TODO: Why does termination checking fail here? - {-# TERMINATING #-} - show-FSTNode : FSTNode A → Lines - show-FSTNode (node a children) = do - > show-A a - indent 2 (show-FST children) - - show-FST : FST A → Lines - show-FST fst = lines (map show-FSTNode fst) - - show-Feature : Feature N A → Lines - show-Feature feature = do - > show-N (name feature) <+> "∷" - indent 2 (show-FST (impl feature)) + open import Data.String using (String; _<+>_) + open import Show.Lines + + module Show {N : 𝔽} (show-N : N → String) (show-A : A → String) where + mutual + -- TODO: Why does termination checking fail here? + {-# TERMINATING #-} + show-FSTNode : FSTNode → Lines + show-FSTNode (node a children _) = do + > show-A a + indent 2 (show-FST children) + + show-FST : FST → Lines + show-FST fst = lines (map show-FSTNode fst) + + show-Feature : Feature N → Lines + show-Feature feature = do + > show-N (name feature) <+> "∷" + indent 2 (show-FST (impl feature)) From 26ab25db67d55517565662009402b7087d2f6d14 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Mon, 6 Nov 2023 14:35:23 +0100 Subject: [PATCH 15/43] fst renamings --- .../V2/Experiment/FST-Experiments.agda | 12 ++-- src/Framework/V2/Lang/FST.agda | 63 +++++++++---------- 2 files changed, 36 insertions(+), 39 deletions(-) diff --git a/src/Framework/V2/Experiment/FST-Experiments.agda b/src/Framework/V2/Experiment/FST-Experiments.agda index 2e5ad3a8..e61aac48 100644 --- a/src/Framework/V2/Experiment/FST-Experiments.agda +++ b/src/Framework/V2/Experiment/FST-Experiments.agda @@ -32,7 +32,7 @@ module _ {A : 𝔸} (_≟_ : DecidableEquality A) where → (N → String) → (A → String) → List (Conf N) - → Experiment (FSF N) + → Experiment (SPL N) getName (exp _ _ _) = "Configure FST example" get (exp show-N show-A configs) (example-name ≔ forest) = let open FSTDefsA.Show show-N show-A @@ -47,7 +47,7 @@ module _ {A : 𝔸} (_≟_ : DecidableEquality A) where linebreak > "⟦ e ⟧" <+> cstr <+> "=" indent 2 do - show-FST (⟦ forest ⟧ c) + show-FSF (⟦ forest ⟧ c) pick-all : ∀ {N} → Conf N pick-all _ = true @@ -103,16 +103,16 @@ module Java where ---- Example SPLs - ex-Add-Sub : Example (FSF ASTNode) + ex-Add-Sub : Example (SPL ASTNode) ex-Add-Sub = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ [] - ex-Sub-Add : Example (FSF ASTNode) + ex-Sub-Add : Example (SPL ASTNode) ex-Sub-Add = "sub-add" ≔ feature-Sub ∷ feature-Add ∷ [] - ex-Add-Sub-Log : Example (FSF ASTNode) + ex-Add-Sub-Log : Example (SPL ASTNode) ex-Add-Sub-Log = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ feature-Log ∷ [] - ex-all : List (Example (FSF ASTNode)) + ex-all : List (Example (SPL ASTNode)) ex-all = ex-Add-Sub ∷ ex-Sub-Add ∷ ex-Add-Sub-Log ∷ [] ---- Experiments diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 798d90b1..64b4796a 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -18,78 +18,75 @@ open import Framework.V2.Annotation.Name using (Name) open import Framework.V2.Constructs.Artifact open import Framework.V2.Lang.FeatureAlgebra --- data FSTNode : 𝕍 where +-- data FST : 𝕍 where -- -- add a proof that the children are disjoint wrt. ≡ --- node : ∀ {A} → A → List (FSTNode A) → FSTNode A +-- node : ∀ {A} → A → List (FST A) → FST A Conf : (N : 𝔽) → Set Conf N = Config N Bool module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where - data FSTNode : Set - different : Rel FSTNode 0ℓ + data FST : Set + different : Rel FST 0ℓ - data FSTNode where - node : A → (children : List FSTNode) → AllPairs different children → FSTNode + data FST where + node : A → (children : List FST) → AllPairs different children → FST different (node a _ _) (node b _ _) = False (a ≟ b) - -- All FSTs have the same implicit root. - -- So an FST is just a list of children, implicitly grouped below - -- an imaginary unique root. - FST : Set - FST = List FSTNode + -- Feature Structure Forest + FSF : Set + FSF = List FST infixr 4 _::_ record Feature (N : 𝔽) : Set where constructor _::_ field name : Name N - impl : FST + impl : FSF open Feature public -- the syntax used in the paper for paths infixr 5 _._[_] - _._[_] : A → (cs : List FSTNode) → AllPairs different cs → FST + _._[_] : A → (cs : List FST) → AllPairs different cs → FSF a . cs [ d ] = node a cs d ∷ [] -- helper function when branching in paths - branches : List (List FSTNode) → List FSTNode + branches : List (List FST) → List FST branches = concat - -- Feature Structure Forest - FSF : (N : 𝔽) → Set --𝔼 - FSF N = List (Feature N) + SPL : (N : 𝔽) → Set --𝔼 + SPL N = List (Feature N) - select : ∀ {N} → Conf N → FSF N → FSF N + select : ∀ {N} → Conf N → SPL N → SPL N select c = filterᵇ (c ∘ name) - forget-names : ∀ {N} → FSF N → List FST + forget-names : ∀ {N} → SPL N → List FSF forget-names = map impl - names : ∀ {N} → FSF N → List N + names : ∀ {N} → SPL N → List N names = map name open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) open Eq.≡-Reasoning - 𝟘 : FST + 𝟘 : FSF 𝟘 = [] mutual -- TODO: Avoid termination macro. {-# TERMINATING #-} - impose-subtree : FSTNode → List FSTNode → List FSTNode + impose-subtree : FST → List FST → List FST impose-subtree l [] = l ∷ [] impose-subtree (node a as as-unique) (node b bs bs-unique ∷ rs) with a ≟ b ... | yes _ = node b (as ⊕ bs) {!!} ∷ rs ... | no _ = node b bs bs-unique ∷ impose-subtree (node a as as-unique) rs infixr 7 _⊕_ - _⊕_ : FST → FST → FST + _⊕_ : FSF → FSF → FSF l ⊕ r = foldr impose-subtree r l - ⊕-all : List FST → FST + ⊕-all : List FSF → FSF ⊕-all = foldr _⊕_ 𝟘 l-id : LeftIdentity _≡_ 𝟘 _⊕_ @@ -122,10 +119,10 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where cong : Congruent₂ _≡_ _⊕_ cong refl refl = refl - idem : ∀ (i₁ i₂ : FST) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ + idem : ∀ (i₁ i₂ : FSF) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ idem = {!!} - FST-is-FeatureAlgebra : FeatureAlgebra FST _⊕_ 𝟘 + FST-is-FeatureAlgebra : FeatureAlgebra FSF _⊕_ 𝟘 FST-is-FeatureAlgebra = record { monoid = record { isSemigroup = record @@ -142,7 +139,7 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where where open import Data.Product using (_,_) - ⟦_⟧ : ∀ {N : 𝔽} → FSF N → Conf N → FST + ⟦_⟧ : ∀ {N : 𝔽} → SPL N → Conf N → FSF ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features -- We could avoid wrap and unwrap by defining our own intermediate tree structure @@ -160,15 +157,15 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where mutual -- TODO: Why does termination checking fail here? {-# TERMINATING #-} - show-FSTNode : FSTNode → Lines - show-FSTNode (node a children _) = do + show-FST : FST → Lines + show-FST (node a children _) = do > show-A a - indent 2 (show-FST children) + indent 2 (show-FSF children) - show-FST : FST → Lines - show-FST fst = lines (map show-FSTNode fst) + show-FSF : FSF → Lines + show-FSF fst = lines (map show-FST fst) show-Feature : Feature N → Lines show-Feature feature = do > show-N (name feature) <+> "∷" - indent 2 (show-FST (impl feature)) + indent 2 (show-FSF (impl feature)) From fbdd8fb89c29d5b77144ed979cf64954d3f503bf Mon Sep 17 00:00:00 2001 From: pmbittner Date: Mon, 6 Nov 2023 14:36:59 +0100 Subject: [PATCH 16/43] remove deprecated code --- src/Framework/V2/Lang/FST.agda | 4 ---- 1 file changed, 4 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 64b4796a..579ba171 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -18,10 +18,6 @@ open import Framework.V2.Annotation.Name using (Name) open import Framework.V2.Constructs.Artifact open import Framework.V2.Lang.FeatureAlgebra --- data FST : 𝕍 where --- -- add a proof that the children are disjoint wrt. ≡ --- node : ∀ {A} → A → List (FST A) → FST A - Conf : (N : 𝔽) → Set Conf N = Config N Bool From 654346fb078fc908ddf342e378c62bc8e3e0cef1 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Mon, 6 Nov 2023 17:35:20 +0100 Subject: [PATCH 17/43] intermediate --- src/Framework/V2/Lang/FST.agda | 159 +++++++++++++++++++++++++++++---- 1 file changed, 140 insertions(+), 19 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 579ba171..ca661ac3 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -4,10 +4,13 @@ module Framework.V2.Lang.FST where open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) -open import Data.List.Relation.Unary.AllPairs using (AllPairs) +open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) +open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Data.Unit using (tt) open import Function using (_∘_) open import Level using (0ℓ) +open import Relation.Nullary.Negation using (¬_) open import Relation.Nullary.Decidable using (yes; no; False) open import Relation.Binary using (DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) @@ -23,18 +26,24 @@ Conf N = Config N Bool module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where data FST : Set + record FSF : Set different : Rel FST 0ℓ data FST where - node : A → (children : List FST) → AllPairs different children → FST + node : A → FSF → FST - different (node a _ _) (node b _ _) = False (a ≟ b) + different (node a _) (node b _) = False (a ≟ b) -- Feature Structure Forest - FSF : Set - FSF = List FST + infix 4 ⟪_,_⟫ + record FSF where + inductive + constructor ⟪_,_⟫ + field + roots : List FST + no-duplicates : AllPairs different roots - infixr 4 _::_ + infixr 3 _::_ record Feature (N : 𝔽) : Set where constructor _::_ field @@ -45,7 +54,7 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where -- the syntax used in the paper for paths infixr 5 _._[_] _._[_] : A → (cs : List FST) → AllPairs different cs → FSF - a . cs [ d ] = node a cs d ∷ [] + a . cs [ d ] = ⟪ node a ⟪ cs , d ⟫ ∷ [] , [] ∷ [] ⟫ -- helper function when branching in paths branches : List (List FST) → List FST @@ -63,24 +72,137 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where names : ∀ {N} → SPL N → List N names = map name + map-different : ∀ {b xs} → + ∀ (ys : FSF) (z : FST) + → different (node b xs) z + → different (node b ys) z + map-different {b} _ (node z _) l with z ≟ b + ... | yes _ = l + ... | no _ = l + + map-all-different : ∀ {b cs cs' xs} + → All (different (node b cs )) xs + → All (different (node b cs')) xs + map-all-different [] = [] + map-all-different {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-different cs' x px ∷ map-all-different pxs + open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) open Eq.≡-Reasoning 𝟘 : FSF - 𝟘 = [] + 𝟘 = ⟪ [] , [] ⟫ mutual + data PlainFST : Set where + pnode : A → List PlainFST → PlainFST + + pdifferent : Rel PlainFST 0ℓ + pdifferent (pnode a _) (pnode b _) = False (a ≟ b) + + infix 4 _+_⟶_ + data _+_⟶_ : PlainFST → List (PlainFST) → List (PlainFST) → Set where + base : ∀ {l} + ---------- + → l + [] ⟶ l ∷ [] + + merge : ∀ {a as bs rs cs} + → as + bs ↝ cs + → pnode a as + pnode a bs ∷ rs ⟶ pnode a cs ∷ rs + + skip : ∀ {a as b bs rs cs} + → ¬ (a ≡ b) + → pnode a as + rs ⟶ cs + → pnode a as + pnode b bs ∷ rs ⟶ pnode b bs ∷ cs + + infix 4 _+_↝_ + data _+_↝_ : List PlainFST → List PlainFST → List PlainFST → Set where + impose-nothing : ∀ {rs} + → [] + rs ↝ rs + + impose-step : ∀ {l ls rs e e'} + → l + rs ⟶ e + → ls + e ↝ e' + ---------------- + → l ∷ ls + rs ↝ e' + + Unique : List PlainFST → Set + Unique = AllPairs pdifferent + + ↝-deterministic : ∀ {fs rs e e'} + → fs + rs ↝ e + → fs + rs ↝ e' + → e ≡ e' + ↝-deterministic impose-nothing impose-nothing = refl + ↝-deterministic (impose-step ⟶x ↝x) (impose-step ⟶y ↝y) rewrite ⟶-deterministic ⟶x ⟶y | ↝-deterministic ↝x ↝y = refl + + open import Data.Empty using (⊥; ⊥-elim) + ⟶-deterministic : ∀ {f rs e e'} + → f + rs ⟶ e + → f + rs ⟶ e' + → e ≡ e' + ⟶-deterministic base base = refl + ⟶-deterministic (merge x) (merge y) rewrite ↝-deterministic x y = refl + ⟶-deterministic (merge x) (skip a≠a y) = ⊥-elim (a≠a refl) + ⟶-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) + ⟶-deterministic (skip neq x) (skip neq' y) rewrite ⟶-deterministic x y = refl + + open import Data.Product using (∃-syntax; _,_) + ↝-total : ∀ (ls rs : List PlainFST) → ∃[ e ] (ls + rs ↝ e) + ↝-total [] rs = rs , impose-nothing + ↝-total (l ∷ ls) rs = {!!} , {!!} + + ⟶-total : ∀ (l : PlainFST) (rs : List PlainFST) → ∃[ e ] (l + rs ⟶ e) + ⟶-total l [] = {!!} + ⟶-total l (x ∷ rs) = {!!} + -- TODO: Avoid termination macro. {-# TERMINATING #-} - impose-subtree : FST → List FST → List FST - impose-subtree l [] = l ∷ [] - impose-subtree (node a as as-unique) (node b bs bs-unique ∷ rs) with a ≟ b - ... | yes _ = node b (as ⊕ bs) {!!} ∷ rs - ... | no _ = node b bs bs-unique ∷ impose-subtree (node a as as-unique) rs + impose-subtree : FST → FSF → FSF + impose-subtree l ⟪ [] , no-duplicates ⟫ = ⟪ l ∷ [] , [] ∷ [] ⟫ + impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ node b ⟪ bs , υ-bs ⟫ ∷ rs , υ-b ∷ υ-rs ⟫ with a ≟ b + ... | yes _ = ⟪ node b (⟪ as , υ-as ⟫ ⊕ ⟪ bs , υ-bs ⟫) ∷ rs , map-all-different υ-b ∷ υ-rs ⟫ + ... | no ¬p = + ⟪ node b ⟪ bs , υ-bs ⟫ ∷ r-rec , helpi (different-values a b ⟪ as , υ-as ⟫ ⟪ bs , υ-bs ⟫ ¬p) υ-b ∷ υ-rec ⟫ + where + rec = impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ rs , υ-rs ⟫ + r-rec = FSF.roots rec + υ-rec = FSF.no-duplicates rec + + ¬-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} → ¬ (a ≡ b) → ¬ (b ≡ a) + ¬-sym ¬a≡b b≡a = ¬a≡b (Eq.sym b≡a) + + different-values : ∀ (a b : A) (xs ys : FSF) + → ¬ (a ≡ b) + → different (node a xs) (node b ys) + different-values a b _ _ neq with a ≟ b + ... | yes eq = neq eq + ... | no neq = tt + + open import Data.Empty using (⊥; ⊥-elim) + open import Relation.Nullary.Decidable using (isYes) + different-sym : ∀ {a b} + → different a b + → different b a + different-sym {node a as} {node b bs} neq with b ≟ a + ... | no neq = tt + ... | yes eq = {!!} + + helpi : ∀ {na nb} {xs : List FST} {υ-xs : AllPairs different xs} + → different na nb + → All (different nb) xs + → All (different nb) (FSF.roots (impose-subtree na ⟪ xs , υ-xs ⟫)) + helpi {na} {nb} na≠nb [] = different-sym {na} {nb} na≠nb ∷ [] + helpi {node a as} {node b bs} {x ∷ xs} {υ-x ∷ υ-xs} na≠nb (px ∷ pxs) with a ≟ b + ... | yes _ = {!!} + ... | no _ = {!!} + -- impose-subtree l ⟪ [] , _ ⟫ = l ∷ [] + -- impose-subtree (node a ⟪ as , as-unique ⟫) ⟪ node b ⟪ bs , bs-unique ⟫ ∷ rs , _ ⟫ with a ≟ b + -- ... | yes _ = node b ? ∷ rs + -- ... | no _ = node b ⟪ bs , bs-unique ⟫ ∷ impose-subtree (node a ⟪ as , as-unique ⟫) rs infixr 7 _⊕_ _⊕_ : FSF → FSF → FSF - l ⊕ r = foldr impose-subtree r l + ⟪ l , _ ⟫ ⊕ r = foldr impose-subtree r l ⊕-all : List FSF → FSF ⊕-all = foldr _⊕_ 𝟘 @@ -100,8 +222,7 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where -- So it is normal, right? -- Maybe, the imposition should not be done sequentially but in parallel? r-id : RightIdentity _≡_ 𝟘 _⊕_ - r-id [] = refl - r-id (x ∷ xs) = {!!} + r-id = {!!} -- rewrite r-id xs = -- begin -- impose-subtree x xs @@ -154,12 +275,12 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where -- TODO: Why does termination checking fail here? {-# TERMINATING #-} show-FST : FST → Lines - show-FST (node a children _) = do + show-FST (node a ⟪ children , uniq ⟫) = do > show-A a - indent 2 (show-FSF children) + indent 2 (show-FSF ⟪ children , uniq ⟫) show-FSF : FSF → Lines - show-FSF fst = lines (map show-FST fst) + show-FSF fst = lines (map (show-FST) (FSF.roots fst)) show-Feature : Feature N → Lines show-Feature feature = do From fafee363c252f6cc89d821fa59f7b9af738aa182 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 8 Nov 2023 14:09:05 +0100 Subject: [PATCH 18/43] totality and preservation of uniqueness proven --- src/Framework/V2/Lang/FST.agda | 153 ++++++++++++++++++++++++++++++--- 1 file changed, 140 insertions(+), 13 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index ca661ac3..d4962e18 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -6,12 +6,13 @@ open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Data.Empty using (⊥; ⊥-elim) open import Data.Unit using (tt) open import Function using (_∘_) open import Level using (0ℓ) open import Relation.Nullary.Negation using (¬_) -open import Relation.Nullary.Decidable using (yes; no; False) +open import Relation.Nullary.Decidable using (yes; no; _because_; False) open import Relation.Binary using (DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) @@ -24,6 +25,26 @@ open import Framework.V2.Lang.FeatureAlgebra Conf : (N : 𝔽) → Set Conf N = Config N Bool +≠-sym : ∀ {ℓ} {A : Set ℓ} (a b : A) + → ¬ (a ≡ b) + → ¬ (b ≡ a) +≠-sym a b a≠b refl = a≠b refl + +≠→False : ∀ {ℓ} {A : Set ℓ} {a b : A} + → (_≟_ : DecidableEquality A) + → ¬ (a ≡ b) + → False (a ≟ b) +≠→False {a = a} {b = b} _≟_ a≠b with a ≟ b +... | yes a≡b = ⊥-elim (a≠b a≡b) +... | no _ = tt + +False-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} + → (_≟_ : DecidableEquality A) + → False (a ≟ b) + → False (b ≟ a) +False-sym {a = a} {b = b} _≟_ _ with a ≟ b +... | no ¬p = ≠→False _≟_ (≠-sym a b ¬p) + module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where data FST : Set record FSF : Set @@ -95,13 +116,40 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where mutual data PlainFST : Set where pnode : A → List PlainFST → PlainFST + foo = pnode pdifferent : Rel PlainFST 0ℓ pdifferent (pnode a _) (pnode b _) = False (a ≟ b) + Unique : List PlainFST → Set + Unique = AllPairs pdifferent + + data UFST : Set where + unode : A → (cs : List UFST) → AllPairs udifferent cs → UFST + + udifferent : Rel UFST 0ℓ + udifferent (unode a _ _) (unode b _ _) = False (a ≟ b) + + forget-unique : UFST -> PlainFST + forget-unique (unode a cs _) = pnode a (map forget-unique cs) + + + map-pdifferent : ∀ {b xs} (ys : List PlainFST) (z : PlainFST) + → pdifferent (foo b xs) z + → pdifferent (foo b ys) z + map-pdifferent {b} _ (pnode z _) l with z ≟ b + ... | yes _ = l + ... | no _ = l + + map-all-pdifferent : ∀ {b cs cs' xs} + → All (pdifferent (foo b cs )) xs + → All (pdifferent (foo b cs')) xs + map-all-pdifferent [] = [] + map-all-pdifferent {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-pdifferent cs' x px ∷ map-all-pdifferent pxs + infix 4 _+_⟶_ data _+_⟶_ : PlainFST → List (PlainFST) → List (PlainFST) → Set where - base : ∀ {l} + base : ∀ {l : PlainFST} ---------- → l + [] ⟶ l ∷ [] @@ -120,13 +168,10 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where → [] + rs ↝ rs impose-step : ∀ {l ls rs e e'} - → l + rs ⟶ e - → ls + e ↝ e' + → l + rs ⟶ e' + → ls + e' ↝ e ---------------- - → l ∷ ls + rs ↝ e' - - Unique : List PlainFST → Set - Unique = AllPairs pdifferent + → l ∷ ls + rs ↝ e ↝-deterministic : ∀ {fs rs e e'} → fs + rs ↝ e @@ -135,7 +180,6 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where ↝-deterministic impose-nothing impose-nothing = refl ↝-deterministic (impose-step ⟶x ↝x) (impose-step ⟶y ↝y) rewrite ⟶-deterministic ⟶x ⟶y | ↝-deterministic ↝x ↝y = refl - open import Data.Empty using (⊥; ⊥-elim) ⟶-deterministic : ∀ {f rs e e'} → f + rs ⟶ e → f + rs ⟶ e' @@ -147,13 +191,96 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where ⟶-deterministic (skip neq x) (skip neq' y) rewrite ⟶-deterministic x y = refl open import Data.Product using (∃-syntax; _,_) + ↝-return : ∀ {e ls rs} + → ls + rs ↝ e + → ∃[ e ] (ls + rs ↝ e) + ↝-return {e} ↝e = e , ↝e + + ⟶-return : ∀ {e l rs} + → l + rs ⟶ e + → ∃[ e ] (l + rs ⟶ e) + ⟶-return {e} ⟶e = e , ⟶e + ↝-total : ∀ (ls rs : List PlainFST) → ∃[ e ] (ls + rs ↝ e) - ↝-total [] rs = rs , impose-nothing - ↝-total (l ∷ ls) rs = {!!} , {!!} + ↝-total [] rs = ↝-return impose-nothing + ↝-total (l ∷ ls) rs = + let e' , ⟶e' = ⟶-total l rs + _ , ↝e = ↝-total ls e' + in ↝-return (impose-step ⟶e' ↝e) ⟶-total : ∀ (l : PlainFST) (rs : List PlainFST) → ∃[ e ] (l + rs ⟶ e) - ⟶-total l [] = {!!} - ⟶-total l (x ∷ rs) = {!!} + ⟶-total l [] = ⟶-return base + ⟶-total (pnode a as) (pnode b bs ∷ rs) with a ≟ b + ... | yes refl = + let cs , ↝cs = ↝-total as bs + in ⟶-return (merge ↝cs) + ... | no a≠b = + let cs , ⟶cs = ⟶-total (pnode a as) rs + in ⟶-return (skip a≠b ⟶cs) + + map-Unique-head : ∀ {a as bs rs} + → Unique (foo a as ∷ rs) + → Unique (foo a bs ∷ rs) + map-Unique-head (x ∷ xs) = map-all-pdifferent x ∷ xs + + drop-second-Unique : ∀ {x y zs} + → Unique (x ∷ y ∷ zs) + → Unique (x ∷ zs) + drop-second-Unique ((_ ∷ pxs) ∷ _ ∷ zs) = pxs ∷ zs + + head-Unique : ∀ {x xs} + → Unique (x ∷ xs) + → All (pdifferent x) xs + head-Unique (x ∷ xs) = x + + -- Observation: We can actually generalize this to any All, not just Unique! + -- impose-step : ∀ {l ls rs e e'} + -- → l + rs ⟶ e' + -- → ls + e' ↝ e + -- ---------------- + -- → l ∷ ls + rs ↝ e + impose-nothing-preserves-unique : ∀ {rs e : List PlainFST} + → [] + rs ↝ e + → Unique rs + → Unique e + impose-nothing-preserves-unique impose-nothing u-rs = u-rs + + impose-step-preserves-unique : ∀ {a : A} {as ls rs e : List PlainFST} + → foo a as ∷ ls + rs ↝ e + → Unique as + → Unique ls + → Unique rs + → Unique e + impose-step-preserves-unique {a} {as} {[]} {rs} {e} (impose-step {e' = e'} ⟶e' ↝e) u-as u-ls u-rs = + let u-e' = ⟶-preserves-unique a as rs u-as u-rs e' ⟶e' + u-e = impose-nothing-preserves-unique ↝e u-e' + in u-e + impose-step-preserves-unique {a} {as} {pnode a' as' ∷ ls} {rs} {e} (impose-step {e' = e'} ⟶e' ↝e) u-as u-ls u-rs = + let u-e' = ⟶-preserves-unique a as rs u-as u-rs e' ⟶e' + u-e = {!impose-step-preserves-unique ↝e!} --↝-preserves-unique ls e' e u-ls u-e' ↝e + in u-e + + ⟶-preserves-unique : ∀ (a : A) (ls rs : List PlainFST) + → Unique ls + → Unique rs + → (e : List PlainFST) + → (foo a ls + rs ⟶ e) -- Bug in Agda here: replacing foo by pnode breaks + → Unique e + ⟶-preserves-unique _ _ _ _ _ _ base = [] ∷ [] + ⟶-preserves-unique _ _ _ _ u-rs _ (merge _) = map-Unique-head u-rs + ⟶-preserves-unique a ls (pnode b bs ∷ rs) u-ls (u-r ∷ u-rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) + = induction a≠b (u-r ∷ u-rs) ⟶cs ∷ unique-cs + where + unique-cs = ⟶-preserves-unique a ls rs u-ls u-rs cs ⟶cs + + induction : ∀ {a ls rs cs b bs} + → ¬ (a ≡ b) + → Unique (foo b bs ∷ rs) + → foo a ls + rs ⟶ cs + → All (pdifferent (foo b bs)) cs + induction a≠b _ base = False-sym _≟_ (≠→False _≟_ a≠b) ∷ [] + induction a≠b u-rs (merge _) = False-sym _≟_ (≠→False _≟_ a≠b) ∷ head-Unique (drop-second-Unique u-rs) + induction a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ induction a≠b (u-r ∷ u-rs) ⟶cs -- TODO: Avoid termination macro. {-# TERMINATING #-} From c7f256c50668757fadfc30e53dcd83d3687dcb92 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 8 Nov 2023 15:31:48 +0100 Subject: [PATCH 19/43] uniqueness done recursively woohooo --- src/Framework/V2/Lang/FST.agda | 68 +++++++++++++++++----------------- 1 file changed, 34 insertions(+), 34 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index d4962e18..87784c5f 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -6,6 +6,7 @@ open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Data.Product using (∃-syntax; _×_; _,_; proj₁; proj₂) open import Data.Empty using (⊥; ⊥-elim) open import Data.Unit using (tt) open import Function using (_∘_) @@ -124,6 +125,12 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where Unique : List PlainFST → Set Unique = AllPairs pdifferent + UniqueNode : PlainFST → Set + UniqueNode (pnode _ as) = UniqueR as + + UniqueR : List PlainFST → Set + UniqueR cs = Unique cs × All UniqueNode cs + data UFST : Set where unode : A → (cs : List UFST) → AllPairs udifferent cs → UFST @@ -133,7 +140,6 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where forget-unique : UFST -> PlainFST forget-unique (unode a cs _) = pnode a (map forget-unique cs) - map-pdifferent : ∀ {b xs} (ys : List PlainFST) (z : PlainFST) → pdifferent (foo b xs) z → pdifferent (foo b ys) z @@ -150,16 +156,18 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where infix 4 _+_⟶_ data _+_⟶_ : PlainFST → List (PlainFST) → List (PlainFST) → Set where base : ∀ {l : PlainFST} - ---------- + --------------- → l + [] ⟶ l ∷ [] merge : ∀ {a as bs rs cs} → as + bs ↝ cs + ---------------------------------------------- → pnode a as + pnode a bs ∷ rs ⟶ pnode a cs ∷ rs skip : ∀ {a as b bs rs cs} → ¬ (a ≡ b) → pnode a as + rs ⟶ cs + ---------------------------------------------- → pnode a as + pnode b bs ∷ rs ⟶ pnode b bs ∷ cs infix 4 _+_↝_ @@ -190,7 +198,6 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where ⟶-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) ⟶-deterministic (skip neq x) (skip neq' y) rewrite ⟶-deterministic x y = refl - open import Data.Product using (∃-syntax; _,_) ↝-return : ∀ {e ls rs} → ls + rs ↝ e → ∃[ e ] (ls + rs ↝ e) @@ -234,44 +241,37 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where head-Unique (x ∷ xs) = x -- Observation: We can actually generalize this to any All, not just Unique! - -- impose-step : ∀ {l ls rs e e'} - -- → l + rs ⟶ e' - -- → ls + e' ↝ e - -- ---------------- - -- → l ∷ ls + rs ↝ e impose-nothing-preserves-unique : ∀ {rs e : List PlainFST} → [] + rs ↝ e → Unique rs → Unique e impose-nothing-preserves-unique impose-nothing u-rs = u-rs - impose-step-preserves-unique : ∀ {a : A} {as ls rs e : List PlainFST} - → foo a as ∷ ls + rs ↝ e - → Unique as - → Unique ls - → Unique rs - → Unique e - impose-step-preserves-unique {a} {as} {[]} {rs} {e} (impose-step {e' = e'} ⟶e' ↝e) u-as u-ls u-rs = - let u-e' = ⟶-preserves-unique a as rs u-as u-rs e' ⟶e' - u-e = impose-nothing-preserves-unique ↝e u-e' - in u-e - impose-step-preserves-unique {a} {as} {pnode a' as' ∷ ls} {rs} {e} (impose-step {e' = e'} ⟶e' ↝e) u-as u-ls u-rs = - let u-e' = ⟶-preserves-unique a as rs u-as u-rs e' ⟶e' - u-e = {!impose-step-preserves-unique ↝e!} --↝-preserves-unique ls e' e u-ls u-e' ↝e - in u-e - - ⟶-preserves-unique : ∀ (a : A) (ls rs : List PlainFST) - → Unique ls - → Unique rs - → (e : List PlainFST) - → (foo a ls + rs ⟶ e) -- Bug in Agda here: replacing foo by pnode breaks - → Unique e - ⟶-preserves-unique _ _ _ _ _ _ base = [] ∷ [] - ⟶-preserves-unique _ _ _ _ u-rs _ (merge _) = map-Unique-head u-rs - ⟶-preserves-unique a ls (pnode b bs ∷ rs) u-ls (u-r ∷ u-rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) - = induction a≠b (u-r ∷ u-rs) ⟶cs ∷ unique-cs + ↝-preserves-unique : ∀ {ls rs e : List PlainFST} + → ls + rs ↝ e + → UniqueR ls + → UniqueR rs + → UniqueR e + ↝-preserves-unique impose-nothing ur-ls ur-rs = ur-rs + ↝-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ⟶e' ↝e) (u-l ∷ u-ls , ur-as ∷ ur-ls) ur-rs = + let ur-e' = ⟶-preserves-unique a as rs e' ⟶e' ur-as ur-rs + ur-e = ↝-preserves-unique ↝e (u-ls , ur-ls) ur-e' + in ur-e + + ⟶-preserves-unique : ∀ (a : A) (ls rs : List PlainFST) (e : List PlainFST) + → foo a ls + rs ⟶ e -- Bug in Agda here: replacing foo by pnode breaks + → UniqueR ls + → UniqueR rs + → UniqueR e + ⟶-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , ur-ls ∷ [] + ⟶-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↝e) ur-ls (u-rs , ur-bs ∷ un-rs) + = map-Unique-head u-rs , ↝-preserves-unique ↝e ur-ls ur-bs ∷ un-rs + ⟶-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) + = induction a≠b (u-r ∷ u-rs) ⟶cs ∷ u-cs , ur-bs ∷ un-cs where - unique-cs = ⟶-preserves-unique a ls rs u-ls u-rs cs ⟶cs + ur-cs = ⟶-preserves-unique a ls rs cs ⟶cs u-ls (u-rs , ur-rs) + u-cs = proj₁ ur-cs + un-cs = proj₂ ur-cs induction : ∀ {a ls rs cs b bs} → ¬ (a ≡ b) From 6e1dbb4466d8791a8b5813b3eaa022d8a05d8418 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 8 Nov 2023 16:53:23 +0100 Subject: [PATCH 20/43] cleanup module structure and mutual blocks --- src/Framework/V2/Lang/FST.agda | 648 ++++++++++++++++----------------- 1 file changed, 318 insertions(+), 330 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 87784c5f..5211e9fc 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -5,7 +5,7 @@ module Framework.V2.Lang.FST where open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) -open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_; head) open import Data.Product using (∃-syntax; _×_; _,_; proj₁; proj₂) open import Data.Empty using (⊥; ⊥-elim) open import Data.Unit using (tt) @@ -26,133 +26,33 @@ open import Framework.V2.Lang.FeatureAlgebra Conf : (N : 𝔽) → Set Conf N = Config N Bool -≠-sym : ∀ {ℓ} {A : Set ℓ} (a b : A) - → ¬ (a ≡ b) - → ¬ (b ≡ a) -≠-sym a b a≠b refl = a≠b refl - -≠→False : ∀ {ℓ} {A : Set ℓ} {a b : A} - → (_≟_ : DecidableEquality A) - → ¬ (a ≡ b) - → False (a ≟ b) -≠→False {a = a} {b = b} _≟_ a≠b with a ≟ b -... | yes a≡b = ⊥-elim (a≠b a≡b) -... | no _ = tt - -False-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} - → (_≟_ : DecidableEquality A) - → False (a ≟ b) - → False (b ≟ a) -False-sym {a = a} {b = b} _≟_ _ with a ≟ b -... | no ¬p = ≠→False _≟_ (≠-sym a b ¬p) - -module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where - data FST : Set - record FSF : Set - different : Rel FST 0ℓ - - data FST where - node : A → FSF → FST - - different (node a _) (node b _) = False (a ≟ b) - - -- Feature Structure Forest - infix 4 ⟪_,_⟫ - record FSF where - inductive - constructor ⟪_,_⟫ - field - roots : List FST - no-duplicates : AllPairs different roots - - infixr 3 _::_ - record Feature (N : 𝔽) : Set where - constructor _::_ - field - name : Name N - impl : FSF - open Feature public - --- the syntax used in the paper for paths - infixr 5 _._[_] - _._[_] : A → (cs : List FST) → AllPairs different cs → FSF - a . cs [ d ] = ⟪ node a ⟪ cs , d ⟫ ∷ [] , [] ∷ [] ⟫ - - -- helper function when branching in paths - branches : List (List FST) → List FST - branches = concat - - SPL : (N : 𝔽) → Set --𝔼 - SPL N = List (Feature N) - - select : ∀ {N} → Conf N → SPL N → SPL N - select c = filterᵇ (c ∘ name) - - forget-names : ∀ {N} → SPL N → List FSF - forget-names = map impl - - names : ∀ {N} → SPL N → List N - names = map name - - map-different : ∀ {b xs} → - ∀ (ys : FSF) (z : FST) - → different (node b xs) z - → different (node b ys) z - map-different {b} _ (node z _) l with z ≟ b - ... | yes _ = l - ... | no _ = l - - map-all-different : ∀ {b cs cs' xs} - → All (different (node b cs )) xs - → All (different (node b cs')) xs - map-all-different [] = [] - map-all-different {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-different cs' x px ∷ map-all-different pxs - - open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) - open Eq.≡-Reasoning - - 𝟘 : FSF - 𝟘 = ⟪ [] , [] ⟫ +module TODO-MOVE-TO-AUX-OR-USE-STL where + ≠-sym : ∀ {ℓ} {A : Set ℓ} (a b : A) + → ¬ (a ≡ b) + → ¬ (b ≡ a) + ≠-sym a b a≠b refl = a≠b refl + + ≠→False : ∀ {ℓ} {A : Set ℓ} {a b : A} + → (_≟_ : DecidableEquality A) + → ¬ (a ≡ b) + → False (a ≟ b) + ≠→False {a = a} {b = b} _≟_ a≠b with a ≟ b + ... | yes a≡b = ⊥-elim (a≠b a≡b) + ... | no _ = tt + + False-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} + → (_≟_ : DecidableEquality A) + → False (a ≟ b) + → False (b ≟ a) + False-sym {a = a} {b = b} _≟_ _ with a ≟ b + ... | no ¬p = ≠→False _≟_ (≠-sym a b ¬p) +open TODO-MOVE-TO-AUX-OR-USE-STL + +module Defs {A : 𝔸} where + data PlainFST : Set where + pnode : A → List PlainFST → PlainFST mutual - data PlainFST : Set where - pnode : A → List PlainFST → PlainFST - foo = pnode - - pdifferent : Rel PlainFST 0ℓ - pdifferent (pnode a _) (pnode b _) = False (a ≟ b) - - Unique : List PlainFST → Set - Unique = AllPairs pdifferent - - UniqueNode : PlainFST → Set - UniqueNode (pnode _ as) = UniqueR as - - UniqueR : List PlainFST → Set - UniqueR cs = Unique cs × All UniqueNode cs - - data UFST : Set where - unode : A → (cs : List UFST) → AllPairs udifferent cs → UFST - - udifferent : Rel UFST 0ℓ - udifferent (unode a _ _) (unode b _ _) = False (a ≟ b) - - forget-unique : UFST -> PlainFST - forget-unique (unode a cs _) = pnode a (map forget-unique cs) - - map-pdifferent : ∀ {b xs} (ys : List PlainFST) (z : PlainFST) - → pdifferent (foo b xs) z - → pdifferent (foo b ys) z - map-pdifferent {b} _ (pnode z _) l with z ≟ b - ... | yes _ = l - ... | no _ = l - - map-all-pdifferent : ∀ {b cs cs' xs} - → All (pdifferent (foo b cs )) xs - → All (pdifferent (foo b cs')) xs - map-all-pdifferent [] = [] - map-all-pdifferent {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-pdifferent cs' x px ∷ map-all-pdifferent pxs - infix 4 _+_⟶_ data _+_⟶_ : PlainFST → List (PlainFST) → List (PlainFST) → Set where base : ∀ {l : PlainFST} @@ -170,6 +70,7 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where ---------------------------------------------- → pnode a as + pnode b bs ∷ rs ⟶ pnode b bs ∷ cs + -- This is bascially just a fold on lists. Maybe we can simplify it accordingly. infix 4 _+_↝_ data _+_↝_ : List PlainFST → List PlainFST → List PlainFST → Set where impose-nothing : ∀ {rs} @@ -181,6 +82,7 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where ---------------- → l ∷ ls + rs ↝ e + mutual ↝-deterministic : ∀ {fs rs e e'} → fs + rs ↝ e → fs + rs ↝ e' @@ -198,218 +100,304 @@ module Defs {A : 𝔸} (_≟_ : DecidableEquality A) where ⟶-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) ⟶-deterministic (skip neq x) (skip neq' y) rewrite ⟶-deterministic x y = refl - ↝-return : ∀ {e ls rs} - → ls + rs ↝ e - → ∃[ e ] (ls + rs ↝ e) - ↝-return {e} ↝e = e , ↝e - - ⟶-return : ∀ {e l rs} - → l + rs ⟶ e - → ∃[ e ] (l + rs ⟶ e) - ⟶-return {e} ⟶e = e , ⟶e - - ↝-total : ∀ (ls rs : List PlainFST) → ∃[ e ] (ls + rs ↝ e) - ↝-total [] rs = ↝-return impose-nothing - ↝-total (l ∷ ls) rs = - let e' , ⟶e' = ⟶-total l rs - _ , ↝e = ↝-total ls e' - in ↝-return (impose-step ⟶e' ↝e) - - ⟶-total : ∀ (l : PlainFST) (rs : List PlainFST) → ∃[ e ] (l + rs ⟶ e) - ⟶-total l [] = ⟶-return base - ⟶-total (pnode a as) (pnode b bs ∷ rs) with a ≟ b - ... | yes refl = - let cs , ↝cs = ↝-total as bs - in ⟶-return (merge ↝cs) - ... | no a≠b = - let cs , ⟶cs = ⟶-total (pnode a as) rs - in ⟶-return (skip a≠b ⟶cs) - - map-Unique-head : ∀ {a as bs rs} - → Unique (foo a as ∷ rs) - → Unique (foo a bs ∷ rs) - map-Unique-head (x ∷ xs) = map-all-pdifferent x ∷ xs + ↝-return : ∀ {e ls rs} + → ls + rs ↝ e + → ∃[ e ] (ls + rs ↝ e) + ↝-return {e} ↝e = e , ↝e + + ⟶-return : ∀ {e l rs} + → l + rs ⟶ e + → ∃[ e ] (l + rs ⟶ e) + ⟶-return {e} ⟶e = e , ⟶e + + module Impose (_≟_ : DecidableEquality A) where + data FST : Set + record FSF : Set + different : Rel FST 0ℓ + + data FST where + node : A → FSF → FST + + different (node a _) (node b _) = False (a ≟ b) + + -- Feature Structure Forest + infix 4 ⟪_,_⟫ + record FSF where + inductive + constructor ⟪_,_⟫ + field + roots : List FST + no-duplicates : AllPairs different roots + + infixr 3 _::_ + record Feature (N : 𝔽) : Set where + constructor _::_ + field + name : Name N + impl : FSF + open Feature public + + -- the syntax used in the paper for paths + infixr 5 _._[_] + _._[_] : A → (cs : List FST) → AllPairs different cs → FSF + a . cs [ d ] = ⟪ node a ⟪ cs , d ⟫ ∷ [] , [] ∷ [] ⟫ + + -- helper function when branching in paths + branches : List (List FST) → List FST + branches = concat + + SPL : (N : 𝔽) → Set --𝔼 + SPL N = List (Feature N) + + select : ∀ {N} → Conf N → SPL N → SPL N + select c = filterᵇ (c ∘ name) + + forget-names : ∀ {N} → SPL N → List FSF + forget-names = map impl + + names : ∀ {N} → SPL N → List N + names = map name + + map-different : ∀ {b xs} → + ∀ (ys : FSF) (z : FST) + → different (node b xs) z + → different (node b ys) z + map-different {b} _ (node z _) l with z ≟ b + ... | yes _ = l + ... | no _ = l + + map-all-different : ∀ {b cs cs' xs} + → All (different (node b cs )) xs + → All (different (node b cs')) xs + map-all-different [] = [] + map-all-different {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-different cs' x px ∷ map-all-different pxs + + open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) + open Eq.≡-Reasoning + + 𝟘 : FSF + 𝟘 = ⟪ [] , [] ⟫ + + mutual + {-# TERMINATING #-} + ↝-total : ∀ (ls rs : List PlainFST) → ∃[ e ] (ls + rs ↝ e) + ↝-total [] rs = ↝-return impose-nothing + ↝-total (l ∷ ls) rs = + let e' , ⟶e' = ⟶-total l rs + _ , ↝e = ↝-total ls e' + in ↝-return (impose-step ⟶e' ↝e) + + ⟶-total : ∀ (l : PlainFST) (rs : List PlainFST) → ∃[ e ] (l + rs ⟶ e) + ⟶-total l [] = ⟶-return base + ⟶-total (pnode a as) (pnode b bs ∷ rs) with a ≟ b + ... | yes refl = + let cs , ↝cs = ↝-total as bs + in ⟶-return (merge ↝cs) + ... | no a≠b = + let cs , ⟶cs = ⟶-total (pnode a as) rs + in ⟶-return (skip a≠b ⟶cs) + + pdifferent : Rel PlainFST 0ℓ + pdifferent (pnode a _) (pnode b _) = False (a ≟ b) + + map-pdifferent : ∀ {b xs} (ys : List PlainFST) (z : PlainFST) + → pdifferent (pnode b xs) z + → pdifferent (pnode b ys) z + map-pdifferent {b} _ (pnode z _) l with z ≟ b + ... | yes _ = l + ... | no _ = l + + map-all-pdifferent : ∀ {b cs cs' xs} + → All (pdifferent (pnode b cs )) xs + → All (pdifferent (pnode b cs')) xs + map-all-pdifferent [] = [] + map-all-pdifferent {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-pdifferent cs' x px ∷ map-all-pdifferent pxs + + Unique : List PlainFST → Set + Unique = AllPairs pdifferent + + unique-ignores-children : ∀ {a as bs rs} + → Unique (pnode a as ∷ rs) + → Unique (pnode a bs ∷ rs) + unique-ignores-children (x ∷ xs) = map-all-pdifferent x ∷ xs drop-second-Unique : ∀ {x y zs} → Unique (x ∷ y ∷ zs) → Unique (x ∷ zs) drop-second-Unique ((_ ∷ pxs) ∷ _ ∷ zs) = pxs ∷ zs - head-Unique : ∀ {x xs} - → Unique (x ∷ xs) - → All (pdifferent x) xs - head-Unique (x ∷ xs) = x - - -- Observation: We can actually generalize this to any All, not just Unique! - impose-nothing-preserves-unique : ∀ {rs e : List PlainFST} - → [] + rs ↝ e - → Unique rs - → Unique e - impose-nothing-preserves-unique impose-nothing u-rs = u-rs - - ↝-preserves-unique : ∀ {ls rs e : List PlainFST} - → ls + rs ↝ e - → UniqueR ls - → UniqueR rs - → UniqueR e - ↝-preserves-unique impose-nothing ur-ls ur-rs = ur-rs - ↝-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ⟶e' ↝e) (u-l ∷ u-ls , ur-as ∷ ur-ls) ur-rs = - let ur-e' = ⟶-preserves-unique a as rs e' ⟶e' ur-as ur-rs - ur-e = ↝-preserves-unique ↝e (u-ls , ur-ls) ur-e' - in ur-e - - ⟶-preserves-unique : ∀ (a : A) (ls rs : List PlainFST) (e : List PlainFST) - → foo a ls + rs ⟶ e -- Bug in Agda here: replacing foo by pnode breaks - → UniqueR ls - → UniqueR rs - → UniqueR e - ⟶-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , ur-ls ∷ [] - ⟶-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↝e) ur-ls (u-rs , ur-bs ∷ un-rs) - = map-Unique-head u-rs , ↝-preserves-unique ↝e ur-ls ur-bs ∷ un-rs - ⟶-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) - = induction a≠b (u-r ∷ u-rs) ⟶cs ∷ u-cs , ur-bs ∷ un-cs - where - ur-cs = ⟶-preserves-unique a ls rs cs ⟶cs u-ls (u-rs , ur-rs) - u-cs = proj₁ ur-cs - un-cs = proj₂ ur-cs - - induction : ∀ {a ls rs cs b bs} - → ¬ (a ≡ b) - → Unique (foo b bs ∷ rs) - → foo a ls + rs ⟶ cs - → All (pdifferent (foo b bs)) cs - induction a≠b _ base = False-sym _≟_ (≠→False _≟_ a≠b) ∷ [] - induction a≠b u-rs (merge _) = False-sym _≟_ (≠→False _≟_ a≠b) ∷ head-Unique (drop-second-Unique u-rs) - induction a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ induction a≠b (u-r ∷ u-rs) ⟶cs - - -- TODO: Avoid termination macro. - {-# TERMINATING #-} - impose-subtree : FST → FSF → FSF - impose-subtree l ⟪ [] , no-duplicates ⟫ = ⟪ l ∷ [] , [] ∷ [] ⟫ - impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ node b ⟪ bs , υ-bs ⟫ ∷ rs , υ-b ∷ υ-rs ⟫ with a ≟ b - ... | yes _ = ⟪ node b (⟪ as , υ-as ⟫ ⊕ ⟪ bs , υ-bs ⟫) ∷ rs , map-all-different υ-b ∷ υ-rs ⟫ - ... | no ¬p = - ⟪ node b ⟪ bs , υ-bs ⟫ ∷ r-rec , helpi (different-values a b ⟪ as , υ-as ⟫ ⟪ bs , υ-bs ⟫ ¬p) υ-b ∷ υ-rec ⟫ - where - rec = impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ rs , υ-rs ⟫ - r-rec = FSF.roots rec - υ-rec = FSF.no-duplicates rec - - ¬-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} → ¬ (a ≡ b) → ¬ (b ≡ a) - ¬-sym ¬a≡b b≡a = ¬a≡b (Eq.sym b≡a) - - different-values : ∀ (a b : A) (xs ys : FSF) - → ¬ (a ≡ b) - → different (node a xs) (node b ys) - different-values a b _ _ neq with a ≟ b - ... | yes eq = neq eq - ... | no neq = tt - - open import Data.Empty using (⊥; ⊥-elim) - open import Relation.Nullary.Decidable using (isYes) - different-sym : ∀ {a b} - → different a b - → different b a - different-sym {node a as} {node b bs} neq with b ≟ a - ... | no neq = tt - ... | yes eq = {!!} - - helpi : ∀ {na nb} {xs : List FST} {υ-xs : AllPairs different xs} - → different na nb - → All (different nb) xs - → All (different nb) (FSF.roots (impose-subtree na ⟪ xs , υ-xs ⟫)) - helpi {na} {nb} na≠nb [] = different-sym {na} {nb} na≠nb ∷ [] - helpi {node a as} {node b bs} {x ∷ xs} {υ-x ∷ υ-xs} na≠nb (px ∷ pxs) with a ≟ b - ... | yes _ = {!!} - ... | no _ = {!!} - -- impose-subtree l ⟪ [] , _ ⟫ = l ∷ [] - -- impose-subtree (node a ⟪ as , as-unique ⟫) ⟪ node b ⟪ bs , bs-unique ⟫ ∷ rs , _ ⟫ with a ≟ b - -- ... | yes _ = node b ? ∷ rs - -- ... | no _ = node b ⟪ bs , bs-unique ⟫ ∷ impose-subtree (node a ⟪ as , as-unique ⟫) rs - - infixr 7 _⊕_ - _⊕_ : FSF → FSF → FSF - ⟪ l , _ ⟫ ⊕ r = foldr impose-subtree r l - - ⊕-all : List FSF → FSF - ⊕-all = foldr _⊕_ 𝟘 - - l-id : LeftIdentity _≡_ 𝟘 _⊕_ - l-id _ = refl - - -- This is not satisfied. What did we do wrong? - -- I think the problem is that (x ∷ xs) ⊕ 𝟘 - -- denotes an FST superimposition of x onto xs, recursively, - -- which is not what we want. - -- What happens is that - -- 1.) x gets imposed onto 𝟘 and yields x - -- 2.) the next child in xs gets imposed onto x, potentially mutating x. - -- BOOM - -- TODO: How to fix that? This self-imposition also occurs when the rhs is not 𝟘. - -- So it is normal, right? - -- Maybe, the imposition should not be done sequentially but in parallel? - r-id : RightIdentity _≡_ 𝟘 _⊕_ - r-id = {!!} - -- rewrite r-id xs = - -- begin - -- impose-subtree x xs - -- ≡⟨ {!!} ⟩ - -- x ∷ xs - -- ∎ - - assoc : Associative _≡_ _⊕_ - assoc x y z = {!!} - - cong : Congruent₂ _≡_ _⊕_ - cong refl refl = refl - - idem : ∀ (i₁ i₂ : FSF) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ - idem = {!!} - - FST-is-FeatureAlgebra : FeatureAlgebra FSF _⊕_ 𝟘 - FST-is-FeatureAlgebra = record - { monoid = record - { isSemigroup = record - { isMagma = record - { isEquivalence = Eq.isEquivalence - ; ∙-cong = cong - } - ; assoc = assoc - } - ; identity = l-id , r-id - } - ; distant-idempotence = idem - } - where - open import Data.Product using (_,_) - - ⟦_⟧ : ∀ {N : 𝔽} → SPL N → Conf N → FSF - ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features - - -- We could avoid wrap and unwrap by defining our own intermediate tree structure - -- that does not reuse Artifact constructor. - -- unwrap : Rose A → Artifact A (Rose A) - -- unwrap (artifact a) = a + mutual + data UniqueNode : PlainFST → Set where + unq : ∀ {a cs} → UniqueR cs → UniqueNode (pnode a cs) - -- wrap : Artifact A (Rose A) → Rose A - -- wrap a = artifact a + UniqueR : List PlainFST → Set + UniqueR cs = Unique cs × All UniqueNode cs - open import Data.String using (String; _<+>_) - open import Show.Lines + mutual + ↝-preserves-unique : ∀ {ls rs e : List PlainFST} + → ls + rs ↝ e + → UniqueR ls + → UniqueR rs + → UniqueR e + ↝-preserves-unique impose-nothing ur-ls ur-rs = ur-rs + ↝-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ⟶e' ↝e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = + let ur-e' = ⟶-preserves-unique a as rs e' ⟶e' ur-as ur-rs + ur-e = ↝-preserves-unique ↝e (u-ls , ur-ls) ur-e' + in ur-e + + ⟶-preserves-unique : ∀ (a : A) (ls rs : List PlainFST) (e : List PlainFST) + → pnode a ls + rs ⟶ e + → UniqueR ls + → UniqueR rs + → UniqueR e + ⟶-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] + ⟶-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↝e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) + = unique-ignores-children u-rs , unq (↝-preserves-unique ↝e ur-ls ur-bs) ∷ un-rs + ⟶-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) + = induction a≠b (u-r ∷ u-rs) ⟶cs ∷ u-cs , ur-bs ∷ un-cs + where + ur-cs = ⟶-preserves-unique a ls rs cs ⟶cs u-ls (u-rs , ur-rs) + u-cs = proj₁ ur-cs + un-cs = proj₂ ur-cs + + induction : ∀ {a ls rs cs b bs} + → ¬ (a ≡ b) + → Unique (pnode b bs ∷ rs) + → pnode a ls + rs ⟶ cs + → All (pdifferent (pnode b bs)) cs + induction a≠b _ base = False-sym _≟_ (≠→False _≟_ a≠b) ∷ [] + induction a≠b u-rs (merge _) = False-sym _≟_ (≠→False _≟_ a≠b) ∷ head (drop-second-Unique u-rs) + induction a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ induction a≠b (u-r ∷ u-rs) ⟶cs - module Show {N : 𝔽} (show-N : N → String) (show-A : A → String) where mutual - -- TODO: Why does termination checking fail here? + -- TODO: Avoid termination macro. {-# TERMINATING #-} - show-FST : FST → Lines - show-FST (node a ⟪ children , uniq ⟫) = do - > show-A a - indent 2 (show-FSF ⟪ children , uniq ⟫) - - show-FSF : FSF → Lines - show-FSF fst = lines (map (show-FST) (FSF.roots fst)) - - show-Feature : Feature N → Lines - show-Feature feature = do - > show-N (name feature) <+> "∷" - indent 2 (show-FSF (impl feature)) + impose-subtree : FST → FSF → FSF + impose-subtree l ⟪ [] , no-duplicates ⟫ = ⟪ l ∷ [] , [] ∷ [] ⟫ + impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ node b ⟪ bs , υ-bs ⟫ ∷ rs , υ-b ∷ υ-rs ⟫ with a ≟ b + ... | yes _ = ⟪ node b (⟪ as , υ-as ⟫ ⊕ ⟪ bs , υ-bs ⟫) ∷ rs , map-all-different υ-b ∷ υ-rs ⟫ + ... | no ¬p = + ⟪ node b ⟪ bs , υ-bs ⟫ ∷ r-rec , helpi (different-values a b ⟪ as , υ-as ⟫ ⟪ bs , υ-bs ⟫ ¬p) υ-b ∷ υ-rec ⟫ + where + rec = impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ rs , υ-rs ⟫ + r-rec = FSF.roots rec + υ-rec = FSF.no-duplicates rec + + ¬-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} → ¬ (a ≡ b) → ¬ (b ≡ a) + ¬-sym ¬a≡b b≡a = ¬a≡b (Eq.sym b≡a) + + different-values : ∀ (a b : A) (xs ys : FSF) + → ¬ (a ≡ b) + → different (node a xs) (node b ys) + different-values a b _ _ neq with a ≟ b + ... | yes eq = neq eq + ... | no neq = tt + + open import Data.Empty using (⊥; ⊥-elim) + open import Relation.Nullary.Decidable using (isYes) + different-sym : ∀ {a b} + → different a b + → different b a + different-sym {node a as} {node b bs} neq with b ≟ a + ... | no neq = tt + ... | yes eq = {!!} + + helpi : ∀ {na nb} {xs : List FST} {υ-xs : AllPairs different xs} + → different na nb + → All (different nb) xs + → All (different nb) (FSF.roots (impose-subtree na ⟪ xs , υ-xs ⟫)) + helpi {na} {nb} na≠nb [] = different-sym {na} {nb} na≠nb ∷ [] + helpi {node a as} {node b bs} {x ∷ xs} {υ-x ∷ υ-xs} na≠nb (px ∷ pxs) with a ≟ b + ... | yes _ = {!!} + ... | no _ = {!!} + -- impose-subtree l ⟪ [] , _ ⟫ = l ∷ [] + -- impose-subtree (node a ⟪ as , as-unique ⟫) ⟪ node b ⟪ bs , bs-unique ⟫ ∷ rs , _ ⟫ with a ≟ b + -- ... | yes _ = node b ? ∷ rs + -- ... | no _ = node b ⟪ bs , bs-unique ⟫ ∷ impose-subtree (node a ⟪ as , as-unique ⟫) rs + + infixr 7 _⊕_ + _⊕_ : FSF → FSF → FSF + ⟪ l , _ ⟫ ⊕ r = foldr impose-subtree r l + + ⊕-all : List FSF → FSF + ⊕-all = foldr _⊕_ 𝟘 + + l-id : LeftIdentity _≡_ 𝟘 _⊕_ + l-id _ = refl + + -- This is not satisfied. What did we do wrong? + -- I think the problem is that (x ∷ xs) ⊕ 𝟘 + -- denotes an FST superimposition of x onto xs, recursively, + -- which is not what we want. + -- What happens is that + -- 1.) x gets imposed onto 𝟘 and yields x + -- 2.) the next child in xs gets imposed onto x, potentially mutating x. + -- BOOM + -- TODO: How to fix that? This self-imposition also occurs when the rhs is not 𝟘. + -- So it is normal, right? + -- Maybe, the imposition should not be done sequentially but in parallel? + r-id : RightIdentity _≡_ 𝟘 _⊕_ + r-id = {!!} + -- rewrite r-id xs = + -- begin + -- impose-subtree x xs + -- ≡⟨ {!!} ⟩ + -- x ∷ xs + -- ∎ + + assoc : Associative _≡_ _⊕_ + assoc x y z = {!!} + + cong : Congruent₂ _≡_ _⊕_ + cong refl refl = refl + + idem : ∀ (i₁ i₂ : FSF) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ + idem = {!!} + + FST-is-FeatureAlgebra : FeatureAlgebra FSF _⊕_ 𝟘 + FST-is-FeatureAlgebra = record + { monoid = record + { isSemigroup = record + { isMagma = record + { isEquivalence = Eq.isEquivalence + ; ∙-cong = cong + } + ; assoc = assoc + } + ; identity = l-id , r-id + } + ; distant-idempotence = idem + } + where + open import Data.Product using (_,_) + + ⟦_⟧ : ∀ {N : 𝔽} → SPL N → Conf N → FSF + ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features + + -- We could avoid wrap and unwrap by defining our own intermediate tree structure + -- that does not reuse Artifact constructor. + -- unwrap : Rose A → Artifact A (Rose A) + -- unwrap (artifact a) = a + + -- wrap : Artifact A (Rose A) → Rose A + -- wrap a = artifact a + + open import Data.String using (String; _<+>_) + open import Show.Lines + + module Show {N : 𝔽} (show-N : N → String) (show-A : A → String) where + mutual + -- TODO: Why does termination checking fail here? + {-# TERMINATING #-} + show-FST : FST → Lines + show-FST (node a ⟪ children , uniq ⟫) = do + > show-A a + indent 2 (show-FSF ⟪ children , uniq ⟫) + + show-FSF : FSF → Lines + show-FSF fst = lines (map (show-FST) (FSF.roots fst)) + + show-Feature : Feature N → Lines + show-Feature feature = do + > show-N (name feature) <+> "∷" + indent 2 (show-FSF (impl feature)) From e1c2b239ab88b7de85082e49df6a5f54b2f22741 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 8 Nov 2023 17:23:47 +0100 Subject: [PATCH 21/43] replace old FST version with new one --- .../V2/Experiment/FST-Experiments.agda | 27 +-- src/Framework/V2/Lang/FST.agda | 202 ++++++------------ 2 files changed, 78 insertions(+), 151 deletions(-) diff --git a/src/Framework/V2/Experiment/FST-Experiments.agda b/src/Framework/V2/Experiment/FST-Experiments.agda index e61aac48..bc57c2ab 100644 --- a/src/Framework/V2/Experiment/FST-Experiments.agda +++ b/src/Framework/V2/Experiment/FST-Experiments.agda @@ -20,13 +20,12 @@ open import Show.Lines open import Util.ShowHelpers open import Data.String using (String; _<+>_; _++_) renaming (_≟_ to _≟ˢ_) -open import Framework.V2.Lang.FST as FSTModule -module FSTDefs = FSTModule.Defs +open import Framework.V2.Lang.FST as FSTModule using (Conf) module _ {A : 𝔸} (_≟_ : DecidableEquality A) where - module FSTDefsA = FSTDefs _≟_ - open FSTDefsA - open FSTModule using (Conf) + open FSTModule.Defs {A} + open FSTModule.Defs.Impose _≟_ + module FSTShow = FSTModule.Defs.Impose.Show exp : ∀ {N} → (N → String) @@ -35,7 +34,7 @@ module _ {A : 𝔸} (_≟_ : DecidableEquality A) where → Experiment (SPL N) getName (exp _ _ _) = "Configure FST example" get (exp show-N show-A configs) (example-name ≔ forest) = - let open FSTDefsA.Show show-N show-A + let open FSTShow _≟_ show-N show-A in do > "Expression e has features" @@ -47,7 +46,7 @@ module _ {A : 𝔸} (_≟_ : DecidableEquality A) where linebreak > "⟦ e ⟧" <+> cstr <+> "=" indent 2 do - show-FSF (⟦ forest ⟧ c) + show-FSF (forget-uniqueness (⟦ forest ⟧ c)) pick-all : ∀ {N} → Conf N pick-all _ = true @@ -66,7 +65,8 @@ module Java where _≟-ast_ : DecidableEquality ASTNode _≟-ast_ = _≟ˢ_ - open FSTDefs _≟-ast_ + open FSTModule.Defs {ASTNode} + open FSTModule.Defs.Impose _≟-ast_ module Calculator where fname-Add = "Add" @@ -88,18 +88,19 @@ module Java where open import Data.Unit using (tt) open import Data.List.Relation.Unary.AllPairs using ([]; _∷_) open import Data.List.Relation.Unary.All using ([]; _∷_) + feature-Add : Feature ASTNode - feature-Add = fname-Add :: cls . add . add-ret . [] [ [] ] [ [] ∷ [] ] [ [] ∷ [] ] + feature-Add = fname-Add :: (cls . add . add-ret . [] , ([] ∷ []) , ((unq ([] ∷ [] , (unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ [])) ∷ [])) feature-Sub : Feature ASTNode - feature-Sub = fname-Sub :: cls . sub . sub-ret . [] [ [] ] [ [] ∷ [] ] [ [] ∷ [] ] + feature-Sub = fname-Sub :: (cls . sub . sub-ret . [] , ([] ∷ []) , ((unq ([] ∷ [] , (unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ [])) ∷ [])) feature-Log : Feature ASTNode feature-Log = fname-Log :: cls . branches ( - (add . log . [] [ [] ] [ [] ∷ [] ]) - ∷ (sub . log . [] [ [] ] [ [] ∷ [] ]) - ∷ []) [ (tt ∷ []) ∷ [] ∷ [] ] + (add . log . []) + ∷ (sub . log . []) + ∷ []) , ([] ∷ []) , ((unq (((tt ∷ []) ∷ ([] ∷ [])) , ((unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ ((unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ [])))) ∷ []) ---- Example SPLs diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 5211e9fc..2d9b660d 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -6,7 +6,7 @@ open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_; head) -open import Data.Product using (∃-syntax; _×_; _,_; proj₁; proj₂) +open import Data.Product using (Σ; ∃-syntax; _×_; _,_; proj₁; proj₂) open import Data.Empty using (⊥; ⊥-elim) open import Data.Unit using (tt) open import Function using (_∘_) @@ -52,6 +52,15 @@ module Defs {A : 𝔸} where data PlainFST : Set where pnode : A → List PlainFST → PlainFST + -- the syntax used in the paper for paths + infixr 5 _._ + _._ : A → (cs : List PlainFST) → List PlainFST + a . cs = pnode a cs ∷ [] + + -- helper function when branching in paths + branches : List (List PlainFST) → List PlainFST + branches = concat + mutual infix 4 _+_⟶_ data _+_⟶_ : PlainFST → List (PlainFST) → List (PlainFST) → Set where @@ -111,74 +120,9 @@ module Defs {A : 𝔸} where ⟶-return {e} ⟶e = e , ⟶e module Impose (_≟_ : DecidableEquality A) where - data FST : Set - record FSF : Set - different : Rel FST 0ℓ - - data FST where - node : A → FSF → FST - - different (node a _) (node b _) = False (a ≟ b) - - -- Feature Structure Forest - infix 4 ⟪_,_⟫ - record FSF where - inductive - constructor ⟪_,_⟫ - field - roots : List FST - no-duplicates : AllPairs different roots - - infixr 3 _::_ - record Feature (N : 𝔽) : Set where - constructor _::_ - field - name : Name N - impl : FSF - open Feature public - - -- the syntax used in the paper for paths - infixr 5 _._[_] - _._[_] : A → (cs : List FST) → AllPairs different cs → FSF - a . cs [ d ] = ⟪ node a ⟪ cs , d ⟫ ∷ [] , [] ∷ [] ⟫ - - -- helper function when branching in paths - branches : List (List FST) → List FST - branches = concat - - SPL : (N : 𝔽) → Set --𝔼 - SPL N = List (Feature N) - - select : ∀ {N} → Conf N → SPL N → SPL N - select c = filterᵇ (c ∘ name) - - forget-names : ∀ {N} → SPL N → List FSF - forget-names = map impl - - names : ∀ {N} → SPL N → List N - names = map name - - map-different : ∀ {b xs} → - ∀ (ys : FSF) (z : FST) - → different (node b xs) z - → different (node b ys) z - map-different {b} _ (node z _) l with z ≟ b - ... | yes _ = l - ... | no _ = l - - map-all-different : ∀ {b cs cs' xs} - → All (different (node b cs )) xs - → All (different (node b cs')) xs - map-all-different [] = [] - map-all-different {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-different cs' x px ∷ map-all-different pxs - - open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) - open Eq.≡-Reasoning - - 𝟘 : FSF - 𝟘 = ⟪ [] , [] ⟫ mutual + --- TODO: Fix termination checking {-# TERMINATING #-} ↝-total : ∀ (ls rs : List PlainFST) → ∃[ e ] (ls + rs ↝ e) ↝-total [] rs = ↝-return impose-nothing @@ -269,55 +213,48 @@ module Defs {A : 𝔸} where induction a≠b u-rs (merge _) = False-sym _≟_ (≠→False _≟_ a≠b) ∷ head (drop-second-Unique u-rs) induction a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ induction a≠b (u-r ∷ u-rs) ⟶cs - mutual - -- TODO: Avoid termination macro. - {-# TERMINATING #-} - impose-subtree : FST → FSF → FSF - impose-subtree l ⟪ [] , no-duplicates ⟫ = ⟪ l ∷ [] , [] ∷ [] ⟫ - impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ node b ⟪ bs , υ-bs ⟫ ∷ rs , υ-b ∷ υ-rs ⟫ with a ≟ b - ... | yes _ = ⟪ node b (⟪ as , υ-as ⟫ ⊕ ⟪ bs , υ-bs ⟫) ∷ rs , map-all-different υ-b ∷ υ-rs ⟫ - ... | no ¬p = - ⟪ node b ⟪ bs , υ-bs ⟫ ∷ r-rec , helpi (different-values a b ⟪ as , υ-as ⟫ ⟪ bs , υ-bs ⟫ ¬p) υ-b ∷ υ-rec ⟫ - where - rec = impose-subtree (node a ⟪ as , υ-as ⟫) ⟪ rs , υ-rs ⟫ - r-rec = FSF.roots rec - υ-rec = FSF.no-duplicates rec + ---- SPL Stuff ---- - ¬-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} → ¬ (a ≡ b) → ¬ (b ≡ a) - ¬-sym ¬a≡b b≡a = ¬a≡b (Eq.sym b≡a) + -- Feature Structure Forest + FSF : Set + FSF = Σ (List PlainFST) UniqueR - different-values : ∀ (a b : A) (xs ys : FSF) - → ¬ (a ≡ b) - → different (node a xs) (node b ys) - different-values a b _ _ neq with a ≟ b - ... | yes eq = neq eq - ... | no neq = tt - - open import Data.Empty using (⊥; ⊥-elim) - open import Relation.Nullary.Decidable using (isYes) - different-sym : ∀ {a b} - → different a b - → different b a - different-sym {node a as} {node b bs} neq with b ≟ a - ... | no neq = tt - ... | yes eq = {!!} - - helpi : ∀ {na nb} {xs : List FST} {υ-xs : AllPairs different xs} - → different na nb - → All (different nb) xs - → All (different nb) (FSF.roots (impose-subtree na ⟪ xs , υ-xs ⟫)) - helpi {na} {nb} na≠nb [] = different-sym {na} {nb} na≠nb ∷ [] - helpi {node a as} {node b bs} {x ∷ xs} {υ-x ∷ υ-xs} na≠nb (px ∷ pxs) with a ≟ b - ... | yes _ = {!!} - ... | no _ = {!!} - -- impose-subtree l ⟪ [] , _ ⟫ = l ∷ [] - -- impose-subtree (node a ⟪ as , as-unique ⟫) ⟪ node b ⟪ bs , bs-unique ⟫ ∷ rs , _ ⟫ with a ≟ b - -- ... | yes _ = node b ? ∷ rs - -- ... | no _ = node b ⟪ bs , bs-unique ⟫ ∷ impose-subtree (node a ⟪ as , as-unique ⟫) rs - - infixr 7 _⊕_ - _⊕_ : FSF → FSF → FSF - ⟪ l , _ ⟫ ⊕ r = foldr impose-subtree r l + forget-uniqueness : FSF → List PlainFST + forget-uniqueness = proj₁ + + infixr 3 _::_ + record Feature (N : 𝔽) : Set where + constructor _::_ + field + name : Name N + impl : FSF + open Feature public + + SPL : (N : 𝔽) → Set --𝔼 + SPL N = List (Feature N) + + select : ∀ {N} → Conf N → SPL N → SPL N + select c = filterᵇ (c ∘ name) + + forget-names : ∀ {N} → SPL N → List FSF + forget-names = map impl + + names : ∀ {N} → SPL N → List N + names = map name + + ---- Algebra ---- + open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) + open Eq.≡-Reasoning + + 𝟘 : FSF + 𝟘 = [] , [] , [] + + infixr 7 _⊕_ + _⊕_ : FSF → FSF → FSF + (l , u-l) ⊕ (r , u-r) = + let e , ↝e = ↝-total l r + u-e = ↝-preserves-unique ↝e u-l u-r + in e , u-e ⊕-all : List FSF → FSF ⊕-all = foldr _⊕_ 𝟘 @@ -325,25 +262,14 @@ module Defs {A : 𝔸} where l-id : LeftIdentity _≡_ 𝟘 _⊕_ l-id _ = refl - -- This is not satisfied. What did we do wrong? - -- I think the problem is that (x ∷ xs) ⊕ 𝟘 - -- denotes an FST superimposition of x onto xs, recursively, - -- which is not what we want. - -- What happens is that - -- 1.) x gets imposed onto 𝟘 and yields x - -- 2.) the next child in xs gets imposed onto x, potentially mutating x. - -- BOOM - -- TODO: How to fix that? This self-imposition also occurs when the rhs is not 𝟘. - -- So it is normal, right? - -- Maybe, the imposition should not be done sequentially but in parallel? r-id : RightIdentity _≡_ 𝟘 _⊕_ - r-id = {!!} - -- rewrite r-id xs = - -- begin - -- impose-subtree x xs - -- ≡⟨ {!!} ⟩ - -- x ∷ xs - -- ∎ + r-id ([] , [] , []) = refl + r-id (x ∷ xs , u) = + begin + (x ∷ xs , u) ⊕ 𝟘 + ≡⟨ {!!} ⟩ + x ∷ xs , u + ∎ assoc : Associative _≡_ _⊕_ assoc x y z = {!!} @@ -389,15 +315,15 @@ module Defs {A : 𝔸} where mutual -- TODO: Why does termination checking fail here? {-# TERMINATING #-} - show-FST : FST → Lines - show-FST (node a ⟪ children , uniq ⟫) = do + show-FST : PlainFST → Lines + show-FST (pnode a children) = do > show-A a - indent 2 (show-FSF ⟪ children , uniq ⟫) + indent 2 (show-FSF children) - show-FSF : FSF → Lines - show-FSF fst = lines (map (show-FST) (FSF.roots fst)) + show-FSF : List PlainFST → Lines + show-FSF roots = lines (map show-FST roots) show-Feature : Feature N → Lines show-Feature feature = do > show-N (name feature) <+> "∷" - indent 2 (show-FSF (impl feature)) + indent 2 (show-FSF (forget-uniqueness (impl feature))) From b96d10a8127e19eebe8ec7f9de5610870eb19bc5 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 10 Nov 2023 10:37:52 +0100 Subject: [PATCH 22/43] some progress on right identity proof --- src/Framework/V2/Lang/FST.agda | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index 2d9b660d..edde357b 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -264,12 +264,8 @@ module Defs {A : 𝔸} where r-id : RightIdentity _≡_ 𝟘 _⊕_ r-id ([] , [] , []) = refl - r-id (x ∷ xs , u) = - begin - (x ∷ xs , u) ⊕ 𝟘 - ≡⟨ {!!} ⟩ - x ∷ xs , u - ∎ + r-id (.(pnode _ _) ∷ [] , [] ∷ [] , unq x ∷ []) = refl + r-id (x ∷ y ∷ zs , u-x ∷ u-y ∷ u-zs , ur-x ∷ ur-y ∷ ur-zs) = {!!} assoc : Associative _≡_ _⊕_ assoc x y z = {!!} From e835cc7f1a4213d34a3f5797cbabc698f56a659a Mon Sep 17 00:00:00 2001 From: Benjamin Moosherr Date: Wed, 6 Dec 2023 18:35:30 +0100 Subject: [PATCH 23/43] Proof termination for FST tree functions --- src/Framework/V2/Lang/FST.agda | 33 +++++++++++++++++++-------------- 1 file changed, 19 insertions(+), 14 deletions(-) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Framework/V2/Lang/FST.agda index edde357b..29483430 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Framework/V2/Lang/FST.agda @@ -3,7 +3,7 @@ module Framework.V2.Lang.FST where open import Data.Bool using (Bool; true; false; if_then_else_) -open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat) +open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat; reverse) open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_; head) open import Data.Product using (Σ; ∃-syntax; _×_; _,_; proj₁; proj₂) @@ -52,6 +52,12 @@ module Defs {A : 𝔸} where data PlainFST : Set where pnode : A → List PlainFST → PlainFST + PlainFST-induction : {B : Set} → (A → List B → B) → PlainFST → B + PlainFST-induction {B} f n = go n [] where + go : PlainFST → List B → B + go (pnode a []) bs = f a (reverse bs) + go (pnode a (c ∷ cs)) bs = go (pnode a cs) (go c [] ∷ bs) + -- the syntax used in the paper for paths infixr 5 _._ _._ : A → (cs : List PlainFST) → List PlainFST @@ -121,24 +127,25 @@ module Defs {A : 𝔸} where module Impose (_≟_ : DecidableEquality A) where + childs : PlainFST → List PlainFST + childs (pnode a as) = as + mutual - --- TODO: Fix termination checking - {-# TERMINATING #-} ↝-total : ∀ (ls rs : List PlainFST) → ∃[ e ] (ls + rs ↝ e) ↝-total [] rs = ↝-return impose-nothing - ↝-total (l ∷ ls) rs = - let e' , ⟶e' = ⟶-total l rs + ↝-total (pnode a as ∷ ls) rs = + let e' , ⟶e' = ⟶-total (pnode a as) rs (↝-total as) _ , ↝e = ↝-total ls e' in ↝-return (impose-step ⟶e' ↝e) - ⟶-total : ∀ (l : PlainFST) (rs : List PlainFST) → ∃[ e ] (l + rs ⟶ e) - ⟶-total l [] = ⟶-return base - ⟶-total (pnode a as) (pnode b bs ∷ rs) with a ≟ b + ⟶-total : ∀ (l : PlainFST) (rs : List PlainFST) → ((rs' : List PlainFST) → ∃[ e ] (childs l + rs' ↝ e)) → ∃[ e ] (l + rs ⟶ e) + ⟶-total l [] _ = ⟶-return base + ⟶-total (pnode a as) (pnode b bs ∷ rs) ↝-total-as with a ≟ b ... | yes refl = - let cs , ↝cs = ↝-total as bs + let cs , ↝cs = ↝-total-as bs in ⟶-return (merge ↝cs) ... | no a≠b = - let cs , ⟶cs = ⟶-total (pnode a as) rs + let cs , ⟶cs = ⟶-total (pnode a as) rs ↝-total-as in ⟶-return (skip a≠b ⟶cs) pdifferent : Rel PlainFST 0ℓ @@ -309,12 +316,10 @@ module Defs {A : 𝔸} where module Show {N : 𝔽} (show-N : N → String) (show-A : A → String) where mutual - -- TODO: Why does termination checking fail here? - {-# TERMINATING #-} show-FST : PlainFST → Lines - show-FST (pnode a children) = do + show-FST = PlainFST-induction λ a children → do > show-A a - indent 2 (show-FSF children) + indent 2 (lines children) show-FSF : List PlainFST → Lines show-FSF roots = lines (map show-FST roots) From 2383d4b4401af7d5ef01608cd067137caa12bf9c Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 13 Mar 2024 15:18:49 +0100 Subject: [PATCH 24/43] dropme --- res/taxonomy.jpg | Bin 79293 -> 0 bytes 1 file changed, 0 insertions(+), 0 deletions(-) delete mode 100644 res/taxonomy.jpg diff --git a/res/taxonomy.jpg b/res/taxonomy.jpg deleted file mode 100644 index 0614d2b1bd447b4de329bdc5d8f430cf9d870962..0000000000000000000000000000000000000000 GIT binary patch literal 0 HcmV?d00001 literal 79293 zcmb@tcT`i+w=No52ndn{A`nnW2rZy=6h#Ojp$U>eK(Nr8fQ4>BYQPZbO+b(kP$Hlx zSWs!wdyy^(iu9s%0Ri9kcitK2{&C-UE2vzni18 z-_8G)|2JFUH?{^rz&8lgcl>{I{QqwkkBh6HGjP%h@bUKrZVm#mxd5_|`+v!|{>e`N zCCC1geJK|Dz&Q*+mUI7KWT*c{4)VY051iBT?|kQ*zW?MgK-RwI9q{j7|JuJE^SNHP zvIMp^z(*Y92eJSeg7AR+f3^LOK5sBVpi`S55J%$w>~qQlf!;-dKq6oMXCHKUadL8jdAY$r|5jIXPE-)9EkDHqh2@`-J z|Bu_>IuH!PQOgD5U_*o0VQd^Qw!aO)0{}+-U+w=@0UL-N%*DyU4FQe{|L4em>*haO zHg=A`UqE~uz-dqpz}VRi5)+QAz^q#ZAq^@Da5*I1TFTnlnqp*K=7}tM%7`vXMn;Fz zv7RiJfv#xit!S(XmL9?KegewkVWatR#U<_i{<+kt8}k=L`Yp$wwD+fHiC%TJ^6H zU4^YDd%t?^+x8!*T+`K67#`@(or7DNA<<7vm=4)2FJf--&cbD@YY>uhnm&wT1(&wF zcZUa0CdrEoYAM{6Vh>4eg!X7;4XoaaW(e&G)$Zme9VZZiSm&-KSL=ZRJY!m+O{_4JSqboijh;f9wT0 zA$!+JxNElS(_Scy^O4c<^{GRB&hF+1Bpcl~D7I z8(OFSg4~z4F2q-vpS^on|H-uN()+)lp&vRreNn%MTF>e}?k>$(s4)()4@y|d*Ws>d zmq_}3aOci3VWrQH=BjIEW$x$-6r0zNegEUpOWyV!e>Z!hDp-7Hc~wXEqx**8j>Kk1 z#dtS;{`j8>wWjlB+d6aW!RCcCFJ8@t3O;LI{|oY;*9uK3vokC`aNd{JIz0Md8A8hv z&$BphL1l_B$}kO3F_aV{9G@7WNsi4ja#H`U{=?|KS)2xhXk{9#EAd+5evd`PIDAT@ ze_SaNFX1 zXxd*8ZsV;_&b?jnge%5Em``LYiaxF$ilGbf7%+w9(aO9;(li$bflp?Mne@j>Q5jE4 zmFzVY8c@k`$E7Ow^x{crL^4{RGm=cPK%lxeVA~-ZlqjN+kOcx)*N<%HtW`!rBA)Tm za2~BC1CSXLl%xeT8YSl>5-mkWu#+iC`p}Vwq$tVMRB2p;GEKUqiE+|ZOof6$LXtJV ztIy~pyB4Ww=o5`0mK+>r{G5%rek%k~*~FE}dd5r7qsn0EOcoMNBNkz4dZN7P8c8Bx zHg-It2+!!E)3Jc$Btp#t?q`(8=q1`}@VdfrmkTUjSYR=NxU{4oUNXM{o!=F>7{in^ z$QyKCd}4xTCBtJ-MR_rBRu(Z-|5g25_h zbP|+FF*s?25J8zjhBlBIG@Z#j1~d(51VxZ@+9VHUV}fBt!=t~*u{7W+hf*D};P7$= z>rF<#;TUf{?i1QdsL7h;P+;)yVV+`nQn377Mieqb9U5I0p+Zsy6snFu=&7b5j41_E z$vo`7`XH(ayHzTd7B0s`@-kROSbnG{iUz0^9feJ&^$>Xo($Xf4#sp84oFj%`NJU5= zlISV;v&cqILXH}oDWavzfEUf7pN^NaFB}rOYRVa)>qga6&P!xU+f(jpxP-Cif1giiTZD_K5 zE3(Cu9U3N|qzG1(q_Fc4P!tl9mMlV#Ae7{T3anBBHLYOAlG1`1CKzx8fu#G6dB-Lp;GMv|RozEP4Jjb_P^RV3BD(jdZb5-En z!b(-IoDm#GMPb1atRgs*#Zt$;H$h=oA_;ae7x8%oc^HoZ895VtMv*)npITDSLr{)# z3?hqurU<5bf@yeGGMf6N2phpn#K`bgY|iUcpPP9xxEgA#(=9MD^?WHIVci_JUE{Z< zZST=BjcRq)%bmRvy4qCp%zvV*y5Q$eCC$M1KhhTdf^_`cU*+m9=#=Wu>|%_L3BNyZ zz7ndpH}p}f=5F}B*YxssNS)53s5ybct<`%jvwsApCVGU2eOpP=Li#WQ5C%MDl&Y_*naxf8-YoQv8lLqKb5V}+7TOc5Am zj5?n_ybdW>#KX&kdceb((2B=HaWC3tbU0?Wt8^l&b;Ndc%pXlHX!H1d+YQON74o~Z zt7Ai}N3_R5e|GZBd}#V#5PrX`#x%B(i-?=x9+NsIJN@Ub=3kKRL71vfmC~qq?&9a$ z0W)_Fl707$#ot*1p^1oV(1#JB4W4XrPI6APWOYcb3ZLOy_19659*-=R6uH$ud*b3d z)X-nhDIgyPyBFjt)l7$++x5SE&}|>H5d7oGevMK;OfOudV@;tyq(*iafdD)iT05f= zJaKwg_*2T~?)%iaV}7#6zM;umvSLTm8VlWgMqc#%P*d`X*Ihh2Q9AhtF2t!LQM`Fr z_eIZ}xF-k>Q5}g&P7Yhstk!gPoOs&wdb7t(b|8{>ukYpFh#B&`7%)De(BG1V*bd`b$eGD zNZQ3)54$#=$v>F)q7{|Nc4Sc2>A=*{yLxiYRK^Di%KoA(!*!2 z1ZO`OY;OR+XA7#I^r)6NQZ<-)F=ogG;ZEsFZCN; zeLWw1AhB}#dU@SmxA41})6OYN6(_c^7y>pu22BsAd5FY>%N5DWQOE`nNFeYXn^f&Y zSfP3N*;VXyyY4_9#x1SM+R@Mw3Cpdemrbm8B%Xb{JB`|2J^QV4)))}=AW{fDGR2B> zr+E90cA-n0o}{{P`^w{wT|4(ZZr1(2z_FW~(7C#DB6t4i?U4M3t(^gt`~XlqP&AbN z6+H2xH0A5I-BPx0)wivI$=%@J->L=9?Q9q(Z29G6XY2aZ2O0R~@x2ss^97gL;u(1? zKYs2X?VB}->KltcjVA#(O7Xp-`zw<6vqq_BkSxNC@#IaV1CiIryB?-xfG&xZCZ%E` z3~`#?Nfc?g2EjhzS7$D_>&h-@CxmCjN9T8Xz}T;<7A2Ph?Jxd&7s_Wwf8Noac!#|m z^2N6*;lX+-0WF*}dDOY>Ky`~7^=9#ya{231U(P7R93KDv7v#BEZRfRlHup^E)o;GX zx;I>Z=$=?RFSfbV7rpm!_`BcNLez$T@6;XLHyEqSxi(6x@qa-qPx!eXyKipP|0bzN z?Ht_C-QKMm=?QBJk*GEo945BrsYtrpa@evRlGh_r5JFZM^&|imt78Zn?D{Z-6|@0} z2e<|xOt2NJFygf0*!>Z1FB1n0nDD0$x;LaA{;u59_;D|XlKrb@_u0;^9M0Rh_ut5G zA56L@FIm=ZZ7C`B9(=|te$zF=Ik9ATBOF*Yxqv_HbRN8ZW6VcHu%jq&1d&XkSj9EB z?}jAp{o3V5eF^)dhiq|oOaUCJ1tL0l4#{!41=%XZ&aQ63&#n(mHR4ahf;n&vM2`p~ zyoeNmPa4_-q%CNwI*eorM{wcar_0#Z$Okze6+WHx9 z!py$?yQX(v#tJDNJxGW{1{MW`;UX+O|2P+s^e+XHDXxs7itCHfKk>{q=%*YPepfX~ zh|8eIYhQIcJCxjRR>78f(zyP011CvC zAaFgVRtSQM8y)%mPrU5gickLaB<27&1wP4?sI@1SG@* z3)Fvla#RtAr*f|;BmJmDS;={{%dP#=IsFV;KdHYq)=EHn(AUroyOtvmHz&)gV?w41 zQ1pT#9Uc7-4}bg#*1d-1eNw69fBxw5>!Gtfx3%ZTE*rkS`$ihqZ`&gz6h%@;@DoYO zcq>?A;F#`BB7!5wg{7F6#7E0ysTuIQs7{roLNP}Ik|j-`(P%#zCK@b-OVlS5BvbWa zjA$ubN`{48w3M|6xtvufhBSHNh$@PLf%AUcJ%L-2w?DM82b3UN*ozS?-%lQ>h=QWa ztl**Ewxd_-H{HLWC|1G`mJJ1}d+h~A)`A7!-ij~yL&q}`S>iA2_z`+mLI_#d=*_Qf z1JWcUyS}<8>r)FJ5N?$%ACM2*xQ1M*#7>JG4K6n_1pg9zW&@nz`h_XPkd2YuLwa~+_32*c@(jp88Lg|`ptot0|MpR zTLQ96&)pcHq!A=(R$K%!UI@VfYA_YLA5|keYJX{cA0wmGyt#d0ae{j{=l$9}M~%d& zsa@vkL_+WOrIx>-OH;eI@3CjFu#qdfr`AhD5urB8jAV3}GvMZBqCM_F^szNXcD~o3 za-$wf=eTf7LxcZ|yuPnZ=(Wxovg|j6Vmd zJE6yNWd%R)*W^vzl}18-eAI4M>K}LA`|3%X%K&69UZWRR|qavs=>lJUp~0BUDt zsWkf^bJ{$1vb;+|*nY2c*EJ-!Mrz0;7KjBx6o@{I1Q)`n2uX`~;|_|FQwb!#+v90U zk#qanRXLXJ0T?Y26Xs5JJm1QjuUBjIoAxVaLZ5`HZ|G%UqPQ zMHUc>o_Ho2IX%sz711LY9sr{E&B-4AOlrJllq!W%!B_qsSmH2l+WR{8{;QN81&`1p z%@{?nsGNAf?l3^5L`CvIq290eYD|5!mZw+GemG%SloGXjm&<-8w|9HJ?8j~m&)FNV z9c3CEu0sFX4gkcqG$j0PuWYxucS9}()^>2`NF*!V30&{?Y8OMN-_aK zJWqZm-~zPjFi)RG1A^mW@e>7lRsqJ zo6fzqANsg-Ln(Ied{Eh}@{nyyz!>t5UH8{cS<7JETR_I-fC&*1Tu_>o5Rk1kpp7}W zULaqY6!9{t4^r(F^8QqDUJvcj)jcyfD?GHPBhgiL;5@J<%-FBCb35N$yY$|Vr7dt9 z9)asBvwyf48QJZhEipUcR-qWwTYL2PHy;W6nso_|1lp4MY!&kE`*G8(&dfK7A^_Nx zSQJS3I#u2H_|mS1xWJW3*=K7Ty5!Y6;<7<}`fLmfstG!hAT5NSiRsX=;2F*DzG|hO z8bK;x9k_pQ2tKwR()Z&|$Lz++J74B zw0ijj5ga(chY%D(5=Y}!aBJU>v^jwpor&N}F9bjAVWvNP?f1K)!o$%PU~B$pAqt=f z2`ZJP&7A?caG%io8(9;xHV1K$JG0B<+@c!+sPFs79^39r3@KzU%JB0fVWFH_#}c;o zFEnrZmw)>;%Ep-9jS0TtmS{|dvtSns@_6~(=;7Hh($JCKF{$qB6xJ)2CnkM$KOuE( zf8bR7c>Ik5$H}`M-_IyLAq#P0X?Q3V%o)L=iNv&#;nI+J0~D17gP26~;v=9e7#O`1 z@w4fLc30;oEkShPZIA(@wqu9VUHx1 z=b-^7WXf6R!FbvbzroHc4Qv%BU6Q+M{fi&Rknolp-T>J(ePMHxS4d5()( zWrn&4v?(a%R8Mq376!=P9wk5WcVrifg^bC2<{2Ko zqf>1IK-BzO_Zo-7dsfeuq|#CWh~U8k!85SD4-|<)L^8&hp94S-JOGi!g_DHnr*utk ze!ety07Raadpg3!9qs#vzW)i$diqBxvak%xR)XbsjYf;80U8SHiqsnP=!!&(HqS?b z*PO)!!SsetIiTLO*eqxUMhb_FeGTWV_23O>ieOZGDEi>CoKc%m7&G~*Cq2<44;!!W zcJhp0#)2FRNUKoaD6$Y7#)*R9d2opXRzx0>CkL4aoQIA`gmNn8?dA6V4*3{*qxq}o z+Cgb?o%ZF8zaZ_2;_X}a)XwCgv8?P@7&J6I2Mku#DwcCU%|~+=6I|~GdKM{D1}E!b zEG3mqtwc!}9(D(f*{=( zL##FXmyBPYO3u`$l1wN+NCPg}V@ky(jv>hVP6FJ6XG5sH&R!1BiSPyl_9t%!QcDhNVivNQtB6fp^JW8=zkdG*3(L7@J3 zS^JKX?%ED8()Q+88w2Kd<5aDP?51QsY25Uo(=CTR#KVCebotBpRb~J~h2~7<6BKpj z;l+!9(KI}R1u!p(z_Xyj!!bY2MEYY4!lf_QTSy%z9q7KJSD$nT{At952{kf~Fm{A6 zXfbMJgq0pN6~oV&*q{N&t5Br4xNszn#MP=?_3iyG=R$!;^R16%E50jKYRJ%Oc#$@J zh_Z)MR~WP_C`*M=7H#IL>YMa0ONq*RMR~K~oc-0c$96&JfG65RS1j_32wvWMO-PMVRUrj{<7*w96_Cr zXl11j>BoZuL^QHEO&|R;6*I631Bf<>U>;~uUYJ87oC6eR3 zqME>sk#v1sZrO!5{kq*%Ql;l;`8M2 z@E9Aqkb_VRiq@k*Afs`q1ktss>!)@AC$InZ@t2|6$wOl+Z^k?^J-~!VFG<}>vYt;L z1`!>{^}6|0MhB&6^*HN3HSYklgb;m`jI70@FA}1e&m3c0C^p7I3h8hb=Jit)J^#HF z(&A}BGKJuADt>flZisO&m#eBqt=HFkQ~SQp-ezdt?FsIt?&8JQujTM$iaHFTmviW_ zCzQ?$r_$xi4rKx!x}4~=HA3TOr5?LT==*XH;r5is1B=R)l>Pjg1rS<#RLPG|dv0S#ub-1R_rUdO#vz z+bvWy{1OpXl57n>2cc3ek|ldAo)}ZCl&+swuo6ne^je?4)Vt-#@`)OAzj#EXE+!hdg@`ina{DL2M-_VL~Z-VYG z9sgbZ)YDCS=Hs*b9=o4Js7|b+r|<$~gh@w_Yix|IMv4rz^<^WYmGz17iP6b}Y^#S6 zErgDzgCp|Qxb&1Q1QCYb*4OD^oUeB>>lr}R6X@_F7G1>jGdm95&qX065lDC#4-p2D zhKA#MwDJJ-DDtzyZZ_8L-S1VM^`%Ptj@`c?!|c)<{gZd<6MyQao!eIfVp8es(3#Ot z-Ge3o5OZe^oBjoz?%p^7A79U0xT+m`#ekKlVN4+nkm~tRCJzVfZbuxBYKHg)gH8)6IBKrKc+#qGDVTmRtPW+8G~lY7op^s*f2i@J;0p6 z5IprhTA4)P(a*w$4H~n_Rm6|3X#c9;ao_s$LTqC}_fqbx#OL6e%d?d_DHC@3j16O5 zzJ1u_-afRl*`#@@W@(K|cv7foT9!pj~70j20>}lyJCxXZBW2 zCER0Y=flFptZq+Sh>p4TjrY^LccxCSJ&0Tm%^lqgICc`BU30B*vru^4Cns6A`=gHI zU0W2xQ7)>(6JEp%hjMc0Wn$Y!Po{v^T!0+af^2~pguzgdCrGvkMuAKOBRMI)w}fI6 zErRQTq_G1CRtO{6XrvEK)N&9tp%q}+*dydPA$T-P4i1ixERr+8dY%v{JGg!0-RkDv zG}VRYiN;j4+HIH+>l(1ptuY$3`u(H3i4VA^MLKk)t9b z@CEcUhU(AtluZ%942*><99ks9Ta?Vh#uf(TlL%TkZ{!I=7?jFOXQRoY0lH2a@;UEA+JV`ZC|r**XM`S$j)!^_;Kft_(X3nqFG)#?3r-S`Z8ai5FfgdJ zm79ohS5!wxx}ddVITFIsz|~oC5OMwT{6q^Y#5qYtBmhi!3))rE!+O|7-=u))pBQ_c zu~RWY<l*#?eO6-|uh;bDrvu`!H{#hEOPOjty?qzRIpNo9&Wlk@=iH3}br zY;ERtM5GfkYs=0n&1qDI;c;=oVi|NeXSm!z#wfK1KaB#Td>^D<#qh)V`1N@#VJu8h zo=A~8DfKvKj~qA*UVoZzmH-ScQsLCNt4f}#7fO$$&xHc!*SJtB76Ak9gtuG>Es@gK zPqXM|&6<>;`o z2dS6q$dXq(yll+~S7E+sk}>MW`jBK`1WWQ^Ok+WP&ZcWFttHVxsVbGvQ%z~_|4~Gy zvyMq@`D$z>skOZp5t1#Y0cmdeWQ#}{7j#5;CzwVkl=I>vLLza~5T+cHJt9mLM`MT= zXrY88jo^wVa&&ZA5}u6#gp!V{Z5|vrv2uz0IHaGLZ_o$0t!^VS-sdLWmF^y8G8}rU@Y?_`DJPr_WTP`4+rKjp`3AFDjF4zrkMzy zX5o$CScrpSR%}*#GRmoi7Dly&Lm2WV$z=mDP9hJN9!VOJ(Y%;3lJPYIVBk)c$ZS%_ zUuF%8^rdI?TEQ_f6uH!hUQo_fwtJz>0(3?p91Ac4W~pcLa$z29E1N0y&pXw-BY&K+Fc)iQf;rq#=GqV-Fc6lBT7Vy%z0?;zU6+ijd)sW~8~(o=@U<(%JJh%h(Y(Wdac0tSEf!xCUFj z&chxZQ&hz~H;99z9GaIAOPfigvy~$_!%Ik1WdPW`qHDNnxM=`7h!Y!{eKVRm8fTFG zlqzHl^UW&ALK-j=2_n&~A`YVanbs7&7%9y%mwp%hBS5L9DWd+*LQ-%qW$s%)mo@(V zVd(hOAc-+Z&$ltCQ;90=Na)|wRa>k+D`C49{9yjEM4|iFgxv3?UgF~gbC3Og+uxiC z^uH0copY?hysd_joiOarEwy>U$d?&lxPCqF4&U~+xyMR~KC9xF&DQ!4!&kRzP@mnS zH;SKD<~}~#SG!+4T$#|ZWsWhwyjr8*U)iwnLhbjvV(qu)rmH)ncZ*~9Q=V%xHqE)~ z&lJY*ou6El8V#1HozDd)ga&WLM_yj?e|K8x+=~nF&EKfWnZF>D&_@T)LkYUsg&TKw zCT2p7L$eAu-tD}S`21$?&+msHZ$+PRr^Sy=O1-OKEnioV^|?7eI%W}=`J-s*u)h4T zc*uoUr`FDJ#HYMGwKO^+?sznMi@S|_)o$S`WMg2Mqj$it|7Q_uR zyb*uWV)HL3S6%~Hm~{#YV&i0E;{pRK$~ZtAz{)I6U||*)2n>b61%(ij$~cn45foZo z3ZtTNQ(8?^PewJ-AdCU5+=2jW71_AhJNK6lS_fcoN}$F)Z(6$=&$0I+`ziA3eU$!I zzOSPx?#9J)%{=-Cxqf4Gn_2hlWK7!2)=lsIyN^;T2<2xD5Rrus{(>$$_s%b@wY7-O z{YeSAZuaNGOj1+JrFWrOV~_S-e|cXc8p=MbGM{%k6MKH1cdJfR&f(*o<-riW1n=j+ z&J&yoY%PA;j|RwTgU?mQ>-D4kA9H+3aJ3wRCHNW)o-KP`zNrfT}H*H3gadntI%I-`Tf2)VZ>@0{={_q#mw_V%#4 zXMlCmCQ$(Z{1;eJbX!JVC+kP4zDEUE{61akUz0o7M{0vE(Ok z>9Fu9p}xGMW%DItO)Z~<>axyuW0Cn!ZC~l!RqQm%D_WNK@-&m_Q_P<%i>a~_bKQ=s zHE+;{3tsdVYwW+ZxNULyqrb5&*chCp_et=EeawOVBcsy z4l-rgf+Dw`))agwqdOiFVm*j{*&Fq?ynOz#wC5q2ktar|zQh%dm+P5!$6?Btj~W7l z13!oN`I;m~=$J;cw#y3;LR*>(*MuNv1!_2gO9D`LR3|eKfB6!^U z@tQn`*`-F#nyL=f4wXTubkVV48w?JMv-(WM!GEmwgM5N`9k_!u_^7 z=cvkS-%85phvXj5zI}D_@^0n#+Vj#z^0Co!TQ{0Dj>g}0x^^k?(V2KBtNG@tUyV7H z-`*W{gm8z=*It!;tNuM*WY$Sv>{yMu7da;CLY*`1Xq5h`CCOWCujoQ%tuG6{^&f3% zh>-Uoi|358rZha<7qEM>uga8hsn*`w^6s}@O6LD|J0d#uWnonUGtiNp?MH4Qk+Oyz zXdb&Vn^UHd4fhf}aq1?gMh`_k6E;sx=1wnGPXa4C9=>IMlD!mJmHCW)CVZF)l`4|W zlqxvJzOr>rl9Zr6*>0=6Vz79x&-xUpr@}13`~xwTWZe-ZkOhL3Tqrw>j*mY1ap-ck z+4aW~!K5L& zDgNq3{s@aR{6kx)F0*m+hJ}950;2I>5ca`dB@ivJS~ayF_wus`HN|Y)f7(^qUOX`^ zHeU2G(fFwHxC;CuY53r=da|aBM$1&&?19de#NOHn%|`HOw9laBxz3W`^fcZxS0!%X zJ}v}b`Xg)fdh^Od2_Gx5?73B=sER`R*fJ~#=Z&d(&nmF>O=^2Hd}vz<0!GX9*IYe# zi}%KuoYRWDTKMb5mfQ5?@Q}~Rp5wB&iS>n1G35lJYO6QKC(Z?R{ZWvtY1_~RxmWAc z(uoz$lE=|W(&vW+p8Xnm)WdCht{E#Az&?5CHum|I*!RPhv-KU9%`|0#PaT$d%vJG4 z`BIDp+=WDB50W}QbtT-pHbWt|w|nsP)H(U#>xIPmE866j4^(Bc@Y@S4(N7n&1i}ZD zbRR3q82k$AFFacP>1ix#NwApSDXO?Jp#0F@Oud|{HgKIcDj()ase@lrRIwaL83-TG zl52^rQsS$4yQOg$TG1KV@K&LKJ>z`rA(Pw2XLW4+P2PT0ykfSJoGPsxsjlS#PB-AL zK(6u~t>6p(m}}{m6r;nS3;W1&DemJQx!=$3-?e2UC1=W$KE=GZ>{kR!s($YWnT|Yn zdDX7s%wawR{sew0_&lQGNElKBnU2W0oHUrJ-5B1VQKS9jXx)QvP;FPjsoofYaGslq zkIwO&d-P!@f7C`TLSXRlg=0?mhA7qhbEbDk#Uz~04YH4@U${X*o*ET$HS0U9{na04 zkYxC+SxVX*=l7)n?`V6bjH><%pp65hBqZh6%vFm&D!Lcta=pg5vt8wixC_n;>VCCH zwzXdODp5MRquNG?(GWeKe9z2;-KEj$CvQ6p+Q#P>#0EvWP?9_3mXQ*$%l*x;LuxOC zPxrnuN2pxK5tnP#-g&LwDDFGTsx{qC1m3RQv^>i4NKVXEr@h&)=W?VwDqj2P>KyI* z6f(E5t@x>dYQ3#(_Ty)ZeoyIDG@XCpuLt&h%?U_-))+YT?L(JX$^5m6oMgXVxv}>v!!}trk9PLe z$^`P?ad#N{aW%|M`c`HSB(vv14uthQLdI=x{C!Q+OP$kmlFu_{9tbkbBHE_=kM`$u z^`8-EH8q%8{*J-+-3i>tSZL2&RWCD6y827sYhik?5M&SV9Zl|CJ_@gG<-cmw*mMx&*?bHks#wI_CpW@EdtUbpJEveB_D!uN&Ra!MP95ZyUQ(lg z_0caBUX!G*n(4f`FZbqE{GuiqXJJSJ~lOc9F-as zWLHiR{&<$6mNRu&Q(VU@a_B(>`te^NYDTob7Ad(E-$ z3N=bl{)u(%>f&mr;ae89D5p7=?r*M$ZL2obz*pKDab!BvSUOMa(`^o4AVa0a_i<8f zyf(uWp2G2p!F(~6i%a#TeF(;p_|=cx33$Uji>sG7HGatdB8)w{Q5~Ub>2~|s`Y=7d zUMCLc!{Sp65`)N~E>a*iwZj{0w}sC>Z&VEYQ4;B9eDBJUGZawfmsgKmR}X!SuTm0o z_d20C)YZMXP<4EF(sH;x_fW&q^{mqCGP%!9<5?nUuME1TpC@Us-V8p3Um(}PmnCfu zd3>)$2c6>;v^nd~5kwDfP{y{fZRIQ!SyTHRZuxi!so{bX2VpAc1sNX>v5L&|tjU>4 zBQI}gz05oS6e9sh+Y}nNg84 zm#@p2KQ^4XTz}t!`t*2dgs-EhKr-I?JhRucIRldvqaZSV=@2XInITW)R2r)N#PSG6 zrIX%&*CIQ7?5MNh-HHmA=aTOl#hlzz(C!n0LZ%VA%Fl^iCQc^BJ z&3U~tr*n)d$NHXUgbSj`BAndsQe4SlW3fR`Bo2tn^Y>}#vhA4jNA=6e_a5mkeiz+8 zJ>D=^HJ^vc@UFsiPy{K;ju8|em2Ju5A5pFb}Xf#%{4!@W!&VBTCvHyEn$ z(mS-F+;?#9d+8g8-O{(7N~2A7Dtk`RjyF+mmfsrdoX<2`M%x{~5@tIJrsrRwngBen zwm>)eZkYM_mfh3cFS9<8Q}6;e*tZlL)U5?ou=SU0q~1-;LEy*Ve?jQ#?&mvSmfo9f zzR`Vm8!hY5izg7uKgkrR4)pugb`7X&Bga%zVv2sL*!~%+z*heSxpptV$(7X3-iWn@ zw=7JHvrH1|&?TyAo`?8b{Z5sNx12LSPTY?E6OsQ*^UA@f{g$2OW(yb9bya+@Wn|7E zY0Uaf3!>+#Y>H{01Vxme)__> z3!7g*uIrtFxF=kH5u0flK|Uj_kS3+&oo5VMbgiP)t>k*~9ZU! z)R_`w2ClS<$E7`H>uonvFZ@t(n*O>pq3_q|j`dmTgjx)XP8zP<-(BCwq(>)YR1y$U%U1n7)v8?!QX<=* zhR5d-7Uc2uHPd~c{O+)#!IcFs_?8r15H7UVwVFu}T zftdQ@xYrzt%c3eA(9g}kN}d~69dTDc6>2kGT7RbnD3*jLim|Eq$nriCkgUt)?QDYt z%(h6AuUtN&s!*0HKP_%ixXi9FmK-7U)jH2>IIR~0qDnPB^fZ|6e}B|J_Cq7cHA%-N z;4c3s8%ym?;m$x$rB{)E?0h?%DBq<|*sY>*&SxaQ`HKCt{Dx&usXBO4eA-89K zjDF|IPjXS9M`s&S(P2Y2-oe+>KE9JJ(7JW0P;f<6L2k@|5%gW-q|{Sq`FSoyp= ztpfL2;tfFRmX4aQxG>VS`sAR(aP8GN#PX0u_ybpB0{YbVh;4T9**7JcpSkTCwK3{D zXE4wQ@+l9mLVbnmKhfpI)a^p=+<1T3&^4{Vps$~hIE0l-(zyiJ)Ho<~IRME?eS6-1 z=jU5ln`g;pXTbDl@iO)&+ubFi0raMmd&H3KC~bWknMSu$EigFbv>jdQae22^I-(^< zM{SC~5;HnRZ+myyq)728(btEg+U%UWgl09}?%IW1UmqD`BXO#1>sh_LS50(P8(0s}#QRSyOUWkG8BnP>>Vr5fTb8>|QA|b|8u5wA~eq?0$1+$~Q=(3XPkDNmoMvw502)8rM}R`_y$sYJjv*znL$ zjIGDUqfd8M-c5$DPf|ioPd08-?K*Qb)@JnsmzEw8yYD|q(GfoCdsdS}?WMvR632Md zpm8#aSmkvc;;+R6gMP-O)_JyANk-(0#wj|LF-4`WWvlX^A{w~pq=o1|xfy!ZbxUaF z)uuCXWY!V=>Pz@@1I>PQBmG>KSpSKZq*q{ZNBxv5VcM&id<9|>r_P=^a~B>XaoRRY z*Vcxw*Qh9V-~#B>YU15PPQGhZ(cC7N;!0O;Sys=BbXZ22-eZ|IPlRQB$PjAXJa^K{ z=6R|P;&o-wLf}p71-BbV`#H{lien?bA)!BB3%}s4c=;_ZA^kGpYXp<$-aO`f?_`wC zmxVef$uk#j4q1kLnUN|{x$uDeeSwCsyoKB3KKbaGU&s>FmL|^oZs^iXR-A;z@oQln zo>F2d10{nW&ppSC9RnW{h}44djl7YBIW4%G1@?77mND{y3Q+df1lT};q4EaU50 zyRmCV#~V1^CJh=oM=Q@7)-Z(qh-EyA+)g*uKd|WO*g=)R_AfVT74o^xqLzQSFaoyJ zdT5{2f1u`Gh|CTSeZd_!vvj%Rw2yqqd%)Nrc;1PMb8dYsDUgCXhU)Vs8**r;{P5~! z#CGg!zmD~OzNw(PtKv!2YqDN_!e@C1`*ed&lY27atG#^u<;X=vjwbbxGSf<$)SYt@ zZwrxd%i*9vyuVdCk)>a4*tk6J>aCJ+hUdW$j|9uvLtp%g>Q8L$jfw7ds7qG5|K__h z;V(_8`f;8rXDf~Qj=*^{r)}ub;iiTp?&5532VB%BGW>x}DDK3)S*>K?>7tFMkyLX~)wQUxEofB7G&l~>zRx?3xd8t^Z_#^AF zcCS$iaru4Drz6?!(`P*gvIAas|0o{ORz%Bal(fD1>i-|MWcj1}%7uQlhgYUT?TvMp zJdQTnjyaWFd^Gw2_!nxrGc0zKDM;s6`4eZPzfA9E}hcB*{)}5bj{(dQx-|C+5 z+0JvjA7wK|FZR=NJC27_R3dP&5?`JjuDd=J2VXyWZa9Dx5tfY zrUDB))JFqU4>&CS?!CbM(rLR0EpWLSR#-9Qr0KSLT!1ik+sp7q(uS zlQp1K8X==wN@}BEf&TF%lgmCen0d?ZUx+&O`2_ zKK!6T=-`6l*X%7n)&6(2b~@Gu$yYewp8$j2J_?DkPoKzr6_uS4_>k}O)Q!yN*Alxu zgV&0GXUAK&{IH9dxOkMK9ou9D5w&nwDi(DV8Khbpiz=4aBxQR)ruxd#OHPrW8^{(P zJSN}sDPs+dP{!{?Mcp@yY@2t6e=5*96VXEFdVtx?PiW53=4w_78eZ;4N>t}D zKs(i-DQFWAtFhW-74^EK@5j$$$8vIOE$cp`(k-InWqg@);034+>4C!;g2%YaSv{=pt=)G4!xDtpP5&M}aq&gsfryS~L%DfSoi?H*C}^%*v? zi#yWsm=VQ}@TduXA^i-3W|yxD_%$^`NnkSh#hadw-k*IPSx;%MB5g`Y55GCS{);p%(1 z2tD_UvB$FJ)6s3qZ*S~D&>ky4i6&&1!|)TU6FbhOh@JJLwqh>KtAxa z%cbAGV|4nKLu?=x{Z#X{wX?`@rXsl`sZm`1&8o`J1A!W^sTa|hV|?C<&VNDSLv5wq ze?fNS(3KtG7~Nf?+#P78-Gcd^tNEmzth-=&+(^^XjIP#;F0b;erI{0c7qs4Yc3t-? z&&{6dy5mk7TM#@=?Yistc+4W-)8Mj2d3J||ljnM74&l-VT6xYfKillN^+(72M6=Dh zWd3xA_`K8Voz+oXaO!#?zUa_ZExs)0_eN^uqubKgN9BPxXZ~&OpYjX5-X%^M`FPT8 zDcztkN8u^O?dEqM4M&&KoCljs!;i3vQrjacqsMMK)g}1w9TZBgX(vTZy^DgKtPXb3 zew}3y`}x+enn{MY^W!a_bC0H*JLVpBW<(We@{L5r_BwW#(NNZQ+3by7h*_&U_(NSDSt215g{CB!@PsjW&nQ}iBJKLQ-sg`XZ zZ~4p8t@)?|>^J6ptA}-9rtOy~r;EZl1=h!(8Rrz)1?JuU66NkGRMgix_W!W;jnS0^ zUArevCbpdkCbn(cwr$(C&53QB6Wg}UxtaIQL?f zRBQfQHdJSFtTt3;a{R6nPi4OzN`9Smu|N9kc=~wy77M>^WMqf{U?31s;BWibKYWgV z2M{6xAw4nyA3vR(yq-QVl7C#^e@tTEKi~y=M3N>mdL^@wE2i!l*BDKZT&U(a7JMQ6 zK$x(K7zyRv86y5xF?y_wOt#aA2cqP1?oWUCjl_E+T||$^akx%&923Fds!|#?s=u@! z8})$;9|u4Z1hDFkxd;+JjDEr-iWem|R3(U$e5V&>-4J!xjslo&k4=fPb>dQ6xv zx=F_Jx{K*@6Dba(Yb(Sz(NpMy*g5pEuiR6EiNX)weqSh~o`kQhKk0NjX(YKw-$L!6 ziw{h}dGvB^I)iijHdwAupNV1EXF#durg6daO=!PQpJEL-G@@jIBm>rck6u~Jkd=^i z!emHi>=Jk6bi%putchr7A??iepH^1JW_1%`tk4FWoTsDK60V8BBE1Buv=UaemOiQ!t?sYpFeR=bj(UKR;2?oH~jW%ojQ6A2iX(zNYvM2GOY0 zf&>T=BGCx`k?8-IzAC?>Gpy!KGMQG3rbx{h#WSSlOcI&W|9=hHgE?Q))YIZObJ|34 zEV2Kx9uV+9%j`Gv{m=2uddLKX-`=!xdPsbB@^QfWe*a-T;D4-#y{g}Y1xGidUuMk` zGrSLrQXg^FNH>7-q-JI$A3_frDHjN>sWbq4Rae`{M2(7t#!98OQr|~Wm!{WSfX-l} zz8|Ju=%%M1-Svb~j}bVNL@j2jyz|Cd>LY2p;2fDYTW2s-KFBzC+e^)9Q9FQ6qh+Sn zL_5v4pSnZGATay^Y*|Gu15S-?A1+|Vc4L3CRua9u^Kv#dhlFv4@$y6%E8p62UPM2A zSrcZcZDwN+<6)cM{(i9n;dHFbSzBvV^h91V`U z(eWJJMx}JvA<~-O!FQxP6v%JRyotJ3F8+gsxM@l{h~JDV4y;eY8_)d$*an zF0qo;^)2!_lZu?^eQU9#+O;KS-dWq5KF1WX6zxIyvWEFC^7fj_RYfh{Q0E8r&H0M% zyl~c`d}jts5h{1#E;0+>E0vW68tH2s-fNV~B=b^XHY=-}oej0yeg6Ip!fdTGcVW+( zgjM5Zy_-a)vF~>v;L(CXvWucE$TE)TpmfF_5pDZwr?@D(%3SHes`P@>F3&jC8jxyq zAUuL(CQxU_dbrZixg4>5MK&04$@!q8iP&W6B5KnxUjZ%K@N4}is(smYoo~BE|HmKe z62kf_;{5j3)}v(!t0)Ye;W{?cb3xYQt82Q}leZOj4_&6EWD<342lMi5Gi$9p?#d!K z=kP<}`1W#{(k^UuRN8Yq_Xe5Qgh2vtY8#YRF;%8wMkH7A3UR~X#fX8?)hHx2eA5Hh zBihey1gBQCkIx0~(yV?QYbC83Th4K39!?Ghz~+478EU|F?z+uQgR_#>Oa^u^ zPm7l!>^#<5kw0w)r%Cl`|AX;Bc7YEw%4yruS2k7J&m7Lz*s2Os7AvZ)%7Vq=^wKbL+zvY3I;mk zJ&K~VL*XK0J>^GTYgd*PYgb1-dh?^P_^y#VOBImQk_2gB=s5`yVg1EpP@%g9GvFn| zu~6E24FkWzl9=dD`Q8n-31`#otFoJT3UwW=nqF0V0TD|rdfZO4o!iKUVd3J~Y}w$j z_=b63(bLjswZsstz^)3&)^JGodv;xB;T9z#(usiShOYLbbbMufL*3n*Ndq>kWgjju zF0BFKXJ2uCx$*9Wt-@=r50W?hlZ#-NFi-<(>~nm)iStKWL&Gs zLpM(#&lr)lL<@(@RFai+3QI;V7!;olx~a-ejLgP)y6;887QKXyh6clyVAkDV$q(Mv z=ope5D%ctNPA`g&vnj6u+Du{(Zu6O7op_BM0wfvyi63N5?msQK zSTIb7q!)L%8|yE%i&}y@%r*}anhz8f=2>bu6J$B3NSn`MNJ|C{?SIfGzKt~Gf+23a z`1`yguFe#3HgMwItZXe73zKB@F}ZvDS0uAt_Ee$T+pS>@?P5P6oxZ+t(>rpIzlto* z4GDiN7DW@jQs_WKh4~h=yWuGO1%SDUBp~-CWKbBS*eVBDJ}9-CGp{<^3gHfC>O070 zx!3$UOsyB14K}RBSniztxg@J@2_H7(OejZy?_NlXH-Sg!z&pLKrZ)R$*sDN3(;YSC zsdukMchbU(Q@hxGozbI0%{(7A7mQF9RAcWi;H79`oq!EHXYUMU%KgV8v9otl@ZGQjL| z{o?n}G!64@&UZ@qh=_GsUhTzt!}ywXEf=htyYvHY)20fD`cGSPZ(?nJ2tqb!a&F<|F#S8S>gVBq@hbS|> zuexhJx!uk8TU@*sd;6DMSL3Ok-p+N|>VikD8&>A2aPX3zxR0JN8|HNqDakG8#WFXE z=3d(>{sd^)(fTryf}tHpkJ0*dw*i_kpR}0@kX!v~e^j0$QmCZTi{cpqSbsv+{Z8IL zIhyi%3}<|3ri=1>(rqrhYMjiYZ^ep9H5Oxw1#dDq1|`kaW%zA1PAS zPA!Wfa5ccMW|ZvEUL01L?@&U)vpx0I?Zrb1y5`DxWc=hKY~4|eew>LYODk?S@oqAM z+-qMtvb&dpC+j*n^ZQgYTHjv0T0A**QzQFvTr$)V58hSK-gfHDjZhY%wzcPG<|^!J zz!4S}ilr)|H_=4p{3j)8`C5#Z>SQJDWd6pzRhGB%+c2K2+t=b!Pzlao02Gr9BXTJE zBT*$6S9G+OQf4}9DGQw7k*v@J(Z`H)^=SmFz2!lR*wPBIHS~d6hj`^mA6DJH;?z+F z*K5&4oA5c@RIBb^z@ubX?`SlTF6q9;T1`N5CXHgXqFSd3si>JU2hh7LwyKjUL*;pT zMG`(ruhHQ5^*sGjy3UgX!pV@h=S@<(O9GipcnxidGul5W7il#o&?U^BwR=-fB{OQ6u zDTX!OUDhnl+VFy*785F&jG z0E<{)v*KNueu=dMkn@uYo$M>DMk}kUP8X ztK1G9=tZz5^8Ie*S{eORr=48wGd4i+mmk3xk7x$z9E?r$7|4N6J$pG!n!>Zed>`vJ zRKg#6e!0HcfR`|b&Dmg6jCze+`@9xgYd?QnmQWkHRh2|WzmMo~68%_Efmk!%wt0s2 zk5x&(vS9XUbQOHSv+@%CgBf~!T9iB@?Xq+o&en)1;w=I3mbGej z*T(aVdvD7>Hqq{&+TLBw<;RaL1ea}ZC_CfQe)rp!OtEX0*hTCqqX0c!|CxpK@B24` zW2DTc)wF6(Y>Z&fEo0<c{&t)4w2}2@ibTjoB~xlE;Z%}* zTgq=762$GwUF)Q!K)k0c9asG9S!#Fd6lBF-4stJgZgMvIvN&tGS4|+>8g3k@S+bZJ zuKSqXn&QEO7^?I}Z81I87#?9VA%jy#xZRa!S{qAb^gN@;qX_j9aI&pV-q0$FY~m;^ z*_Ug*H_DqE{atEa`VG_)DZ`X&+bqvdnHJc`7tn z^EFtK!X&~tVf@V1jPfBe9dTUiX^Dlxx*WK0Rn5Kj458uB(&fv!QUnir8wp(2_~EN@ zRtS8KK9?7HT~2a6W$GIDRXk2Cw04$F9fr(1rm@Z;$iV5shgTXZ?<|Kqheh>1Y`ptk z{x_P-DQ|ZTc>L5d&+nsNPzV8;BzsI-+Dgw~0`-T{$SjgBOL%Vk*do$eXSNDFreH00 z;2r7Pyn-~A8VN8$@vdAxHR^3`g5GHq5Ba7n@|{a*?m0B8)E1Mh8V_1%+`OEfAZEAm zYayK!Shg0O@L7&->VS$s`&0Z|I9ctI94s2V&(DL|4a48F)d&LuWr72N#n7+)0VMhiqO--l}D}$@G*Y_Siu(^Ou!WuYN1p z!Yzw21w|rg*eVlOx1PNAY-E#9>iYipeNNqY?kh87}nuy8~+6$45IG7GS91~9y1#cl<@D>n_z}C zUnii%(oLkVy?5HL?HTCqSBEUgX5=pQ2pyan;s5G{gC)eLe9u@i+kZB${3wlf*Id+? zgZYMP(ij_i;j=HY#ZGxJ9kPr_th)1@7(WYKvsTwnC$RzQ_{^#=@2NPox$GM@5uH;T zej|nsOyk7_3+753Ys!Ak=@XBF>7Z+mSzmAqOcTco7>;rzcj8Wdgc*lzQFu*dj~BF$1Msd`TD%T_~;f4|h4Hr=|_PBZ6;O4++Z568W@!_I5^ zK3s7yKvv&bg5k+a)V^YIYG@$q@C{pZlU26!4(8_NLnLe8bNvqq$JE9{Udz*2G ze{b>c{8X3iOw=tO_z37x4*goGS?Dwc~lKKL>^ale+;%hAXV>xc1;-=Yv3-)aZNr-vZYotU8V7K9)4fk42U() z>w~FW`&DBt2}g+VtLlx%w!Udb^+YTPSx-}8TbNRe!`*|={`X-PVJ6SH1;k)kDoQ;J z!F{NCK8ck*xu@h}N0u*r4dGt^qc#qMH^(wRs`tQ5qCwFT4*LYg#mCcgXa0NLH+W{o zF~6uj(U|qC^y@zFvGD7A+uk(kA3|~(a~GE}Yw8;-fxm!~1wS(pec8>Fc)EdK`I`bT zl))!kG`)?8S>QGe_-soABhD;kCdyNcRrk!O(w`-@FQZK2aO$EZH!p6LEEx%&mnrs* z+qOsw%1DIw8Xbl4I$Mn-vjQw?{$23B4XhF(cjjx-gy{~WqGE2ZqzCQ|BRzVHsd)K5 z3ej7){7D;l=Gh+OQKE}FuTV#(KIM@aiAW8u zF=J{#WHc8;`m|HEe3tX3Mh+boS=kMt3Wu4vTlFD}Kns)nhG~PbX zGh`SzHPyKc_e7;e%C6|6%eg~p?i=iFsJn+#T6uXezko;FKZWDV-dklK=evPgP2uAw z?F3VwoA95Rxgj(*+zI~^_4TN+x1X1nqbU{|z2`50xjDxzVAzw|P7C?FV#GujiCgX_zlYb;>Wq59exP)+q%`=?c$;jzUzFIiVa0RD5sqvDU-~ zRlx+)H``CjJ_jA$oTVzdR-@;pzW}VifOTriNIxGRYoGY>KRq3&ySHiK@ZIe(U(Yp* z-7OX;`h;%^k2Au$4$tYSa0!Jh%2wOKlzRw%yptypNMpOD`zzd=I~Zvxk~UpFu}CA^ z#)CgETc0ApKse#q*IB)jN!_@A3m1YNavoUA4zada(r&BS@lN2ZUZ(#E>y*Y3YApm3 zrXrS?c!@OG<22*0i65faqr}b?Z}?1l=jm2ViAt!AUaoE_jkiHlrJ~i!8fTui{3NR8 z8g47oW(G48wz?bjpcLx_`~B1m1~%KOVC_s<{CXW6XtVWproXFe?Nb+mFZxuooc|?z zsrwcr_`=)ddnWl8VAHPk`|+ofyHc))O4^HT7jKi(D?h}PMEhOZo3Q10uEKKgp}CKr zZ<6^Sk1Q^Fh{Mz$>gzC45(SfoBe={O_M&+Qf(sP7snbT1QnFKe%(DP2+-v z(Mgu^BNGO@3^;h+w75d4W*;1~AH9Sk@w;nZ(rV|@3Vg6~3+l<)yvVZDt)>)q4?&0V zPCm@z-zupk3cpRd&6771=OV64H6&B#60a2&H&N?TF>kw)O?<1;L+|7j(7$IQxcj2% zMMlbnF^%Kvkl!vve>QkpK5VE4`~}d^TPTGY^RM2>9QNb1l208bxi8f9=OelT-DF#e z=zl0Lk!{QG&WJW-t*m8gA)I z(%kJ#@?H8+-G`foQTlb=+I7aR%6IN#2IJ$D!o%4fD?hF6i`D4(EMh{PJKLF6vBc;r z8-=F0n_F@5v@~`F6Q2OJudrJ@yX(EdOO^rBuZyGHYs}eK$V$SKU7NA8kFs0xs@a7( zX11PM{PdyAz^my%i>y6Aj{d{B1^-grlf?tRAQ3evs?*V&x~=vvpzAh4|CJTSD1p&H zBe=Hq(|4OTT5|ci^r5qlQHHWB^tr9ZdhgMHy9PT{>%E9G2t9u+YT3gl-N7K4NR8vn7tcCbs#)BVrf zWo=kpFjnWGGoH7x2-x7%z=(D_`$2qx>8~tazpow);SQqXFM9E-TdPh>lTO<`hdo1e z?r$|_BE2jKK64Z5Ux4C3YNX2|xm7VxY{nWs|Ni3u0|4tB(6Gqd`hL zFO`zj(EMgg;rW}qY)?8e?01o(2yE8BCLw`q)iCN~{M0!=cq9pWWyCeLhj`A0GkatcW_;Cjp%x|KGCh>e< z>EJnH;OoK*vm=Kd&wAI@ZLT_h*}w9|Yek(C##FYv-Gcts9@nDTJ1M8)109*Hii~l$ zz-x~53m^R(af|bo>v*j{|21%zba*7%)H;QkTX|cZ-eN`Fja!FI9Rij0Ry7ClA4yFP zR-bNtlYbhG2$EvyfU1~9e*tZ(5oZ#d=I_cz#7dPp!hxqAvNM0-kL`?wdE>;g)JI}! zR&T`)-Y#EE`&!P8M`Fo%cG$Ha{@z zGUgg8h>aRhaa66|YA=vc-zba)r>@vg z%HuJ!ML#BDM2Gk$kze45UtQBk2SY>$=5Hvyr0<44T4^Bkkmem+;0F)aF4MHx3KyRm z7@fDkn-&pg*@p0$YrE3*giq1eJyD-gN%TKOdsLNTeZqrNU;a+o6TvrfXqMMxzg>wE zDsPMjXZ6g8wOpR4>2CWPSNS!E^{ln-_bVgfQrIRuKGMuo>(dR>bfz`^ z{b*BY0m{J|C-$!yuM)#c&Fk8dwM)z|ZX+$bZhGzA3wuo-cN3^i$4s%n0H+N;%u4W|&&HVZ1{2NX-94&a#LKEi zxk?9nZT5hDB(c7YS4V1-0Ofe^kjCY~#)ek0Wp}1Hy|l2R&wLi}22QFp+qh@o=0Cf{ z+$EI_hpa8q4oeOLKWjYtT4d$X$&AGjGul0`5-Gu+lT(83X$s#Wso%&~vH^iEO##dwQrN0`4 z;(|#RP7|6|Kh||o1V>&yf9oe68AO1fX^Q5TPkSWXEfgAaZO@wdz^W>ZGAasVEL0M{ zqY4j?d;r1k_6Y;#=C@(NFxZWlAw$(_sTd^~;(^sbPsSfox@5S4zdQ2IGWraaHBDmb z>2L#TEBvX1$QXHaoki`W+(EpR!zFd|EMeTBy}SK(r{Y%=J4d?L7|UgIReURvAeKtW!eg@_d3BD+?MMsA8Zt8->1ilwcNv}v&P+T zUMjIJNe7zlZc*du?wg4}DqQ9a+T=kcGKS#ctbB;H9mcIV;Y$4E^*AZfBhoZ@SKn>b z>7laS95t9BxdTFQe&4xAzIAT!(3qd*+R|G;EiICiKGi465^h9DlO$j_)N$_*S;Ju6 zH_icD=Yce5XsL~+5M zY_x2u`R=`+5|iK2Z>LV7{?MoUoT+5w)U_+K^OxibI6hd&ZTt8I!{ld_?m>N7dn%q&Fa4tL9axv|>*&D3TEl%P77=D~LQ!oqjzNY@zDBN}LnqMQ(VL z(TbMLtuLVF>G}>|@Sedv8#OJbyMP-0xmq^6z>#@kHh}Gh=?r*s6p(2_C#o;2pB>cy zz==OE7A33XCdNaDJ**?6ZB-aSH??IHWFer6uo8~uLToih(ORoRe4C|vr@(Cy{T8~t zWn;1u=Gh^+mCNY8(#qxp80r=G-1!IgjCPMA+U9P)P*})lh*+kumGJlIE(imw?xq>X z7BWz&V&%nYb<1V(QCnRlAXu^Qdd2_Mxh zOttzSNs%r%+$eYKxQHHuvYaJLw(8}$Nf8%iAOKet;-dw;GnMZ0j<%{kUO4lbvjv{> zH(|%GalF{$)1#+_6?MCoq0TP$qS*58H@@^7+)xSG4oONIuK7LH$r6^C)$BKz) zwXO)u{=3BgD9Fd)XSLZ&EH(+he%+D0Z!D#meU;C{qu;4SK+I}!sWo4`EdXprDUCD9 zB%X`4pY}ENreU_)zXQA{!t}okPq(ZJp)kJRAj8`dOwQXLB*_h0&{{9@q2d|{ZmL<8 zI+#q7XoXUlvx()YM;)y#5fi;@YLlf9-iT|>S=_QnlQ*VypP?f$3+Z;)tLSqcLw?mE6?bZ1@j%a^uX^b$;t_}fYp zmzkc2qv5{3YHqrCQG+>g>xJyOJ-RY5bhQk&4(Mx(A+Pa)#DUvV1{y0g!uElwywTO< z7Az?y$h?8-Qdw{7I@?XQ80uCponEiMfRS{vAU6v1eh&kt0L#lxf}Zr(CTn*=B>R>8 zqwCU4giI738J=x#12izI3u=8er*XNU4fDN0=-~kkCyoMfW9A6;80~T=+6U!BXZ}m( z9h&oV7ccETT4*1h3}ovJ;#7(z9j0lr>bIPGojvo;NF>CG%x)1*Nnl&6Lt6Y`4Aa|prr@XVeqg;;hiVV6 z%iA^3QklEF_k=AAf>t`;+R46+lYLD#F(O;`7tr8UkfoOM4Ecc@It-}fGi>#+#UI|D z@ACT#sCrTGn!Pn-L;zeeCvZtgZ7@Esqs*IRlU1%*S@w+CUaK6!aMDg_sD>S>Y%ZQ3$+l*Nr+70QJ*ZT$<({dUWXgT`Ps7DQLu&$M|1zrB;JS zJjxHEs?rd(i5+kCytcnOS6CyZ1bIzo#p>u(71t#`Dk97;k3jf|j!B2P97Vt9Gk)@be9udU z_kJK9yB{Yr#N!adghiOW@#Ec`L^F7h*0z1)kP$GD3526<>SH;2)_t#H-uK!*Hv!eo za{WY*XQyUvW5omN3?Y6>+;vA?vqh|bw94@0!i;ySH2+7XB3yO-q9WQd$@}>F$Szcg zh3Sv8ie&8@Z zQ6j~#Sc6#l;8>c%II?YKk%<&Ue-w4et7*YVTi11btVKeEYUR;5S_D2@J}V7{$H0i4 zM?||bh!=%BCvUY0pmVnc>or(Zr^_zOMPId6cqiR-A4@f{pxHFS+gQ~XwBP8fwH0E@ zndP}f>gqlr!5Ddaah_k53=1*B<`|>I>=)CDEfJeZj1}1<_{5$o?qod-JOEWr9Nf|$ z+WHobiDm@lHDxJ9+i7@sVO%l}8R*NZmPO(5SF zd`zT}=-+entG(8=Ei(QC+-=n|bB}}q?3!a<{3w@U&h2s(CS%>mg0p%HYg{nHeA#JN zIhWaz5$Il(Goz2#gm-pRFG+kmWigh0fKRpk8R%xTsBLB5Msjve_e=D6Q!H>wQ`N~z zWIY;3WYRh|n;*fSsM&~GzA$UD9Ecg52K)zLf~--c{iVLpD+Vn;GT{9B$W74+5m~Po29gchLtj{>u8Y?gP9*G{fC9s7Q7wL`cT-1|J zlx+T1RqyhPFL_Mbo$W{_2(6)-ch7b%a_Cl~<;!0{9W=0W9enCZB{(Vmsyg0gPqrCT zbvazXaW2;%M7InGxII^Dt3q#Kl!WnHF>WH-qi}0e<4mWZ zHf9@Ny!}?7LshL+E*6gS%J*m#E8`KnM$G1?D1YZ&1s5B-=h!L?cO_a+-$a4%X|Vqb zFkpJ2Y5xoOu)gAbt=j$zm?(MM);#y+HMXf*Y;RFH+tk)LTiVx*wqb7lSGdel=k4>y zC6qQVb|*_(tXFtD5B=vevaBQT(|#UX-M)Wbwbv3GlTnK=?Y#V2QQdTi)|pq`v#MQs zQZ*^ zD4#W`LOfUHD*iy8_H%7VZAbN@Orl$+Qs{8BsX6`X znLM&8aVCruhP9{Kr4(Toh(9>6?t-GmgwxM+?W0!S$%o%f9mR_(qj$>jjJuM~{2Fr0 zajXfh1XoVdD73eQ1jhD(cLI^7LygWtMLHJ)@G~|0IfOYHBbIjxmwCKH@2S*s3x1r{ zJmu}wKQ$Cjs=61YsWnM96XH~3kP6NKi(fPm^}Jw93G4-sI9lqip64E)%158d4odLO z+6-YY$=jMEi{D8*SJ@Wiq`GBQ+qS-a+t~SQyuN2u1>W(hqSMbCb!TnhXgTRGAgqSs zF91T-}i-e7t&e$ zpWCdkrmXol#D8@2t^Z~aW9RPibMNG(cgjkh@~!`s{@Ypgq^jnv{Qd2BMf2i&bpJ-` zasK^m53$-WAD6Ye+s|2-HK%NEeZO7LDgL>hgMb2oen*k~H_!wK5%3*nLclL?um8`_ zoUU*7W|!~w|AMda5}HQfSRn7QJu&ejzQizNM~Pym9I=hiSBQ!pS>tk>4o8ZBA&XI5?}xzPF^P& zAf15NTd_mLr}7zrjE4L_KY#}B48VJP)kX;pM66x=Ac)XDom}aE(@WKo?N(^?YWyn8 zRT!$Qjh+B$h9c9Wb_DrZfd;e*5;MWgUvI{Oc6$JLz!64{NS!|*hcTj8R&L@F){@Te z1{EQ42^$iIk{2?F(9seyt>@Vl8EISbM~qw}05u-0(??$O5U_#@!I2k|!l2p!fP`3| zUs65YwK>eA+)B-kh$^;Hp#yYyVQlF5W5eGDbsp{+VerW4hp1uk5l2{R9GQ>YVr9sj zw6-ORe`j;@tBOJo76BV!vhRMd7VgszGePRbVLTT0 zj>8I#ni!m>V^Ohp6ne$<9g>wZ;2tsuR4ng)OL#zb_X}=?^Q9P1$7JzQVDFEzt zR&fY?{5W+Pj4ySX(#MNjlgEM*Eo2k{+0pZsgWR#e)(ncqu(E!gp|6AS(pel@}*^W!E!Kqs)`BUtaN zh+%JRt0>(2H*U6MCWVjgW~N%wFmSN zA6xeYX9YV!3Fgy;Il#f={dx$^JBafG`8nVhE(8r*yiTWz;SUjKLr4$=KqtwApyQ6B z5`nysn=q9CTjq2{k^-aKF z68pCZe&MVJd`rm(!&Za<(q439LIl5`S%`7Xlbl+DpcmH9clmkxGg8%#@OE8NdO)p+ z2CNWn`UqCGk%Z9C_%J&V{JrMziin6s<8CGfat7S64<=-ZBmhA!UH2uH>L&?K$a2Rjxef@V2wMH-%`Y6f1AIr;>x~!9t<^g5a2e0N>Tc7lj=Vq!!SJ(@GhJ z(!22ub|wd-3`#;^gV+bZq~06MpAw;4#i$`Y z=pKNOAiN!OT%1rnVyng^(8rG#Bz#SnB;F>_e{=-+2xtuA!rnK7EQ${ElZ7SQJzC9t zgmR$i2HEw>FY*llCS(Y**FS==C+bbiMDOn=Qy9ld3S*T{ti3Cog#e-~k{zURQ;PK6 z-3RYy&>|Z(x&nUcgeGrA;1dIi!)6_xB@($Iq#jB9YLp%yuJw_jZKP!tEOZ5gOdvq5NDU^0O*)O9M)#SSKXlG#5oZFbnhFW zKr+zO&{m*O=K+LTY7KfLKzzSpW&#nKpMJU{^)K?H^8#!@iTi)>AqTL-*m_1DK~M}X z5CDOUTU@gTI0(Rr$sTwf(MW*Mjn1dViBT4}9u5{R#vwak`*k2foy~$atFYjV93ji^ z1n^Ukcp?sLY}#-i;)4z(K~a&TP~8o6#27d+g+;;SJ?+rI>U+yGgxODC82^NceY#7T?@ExUc+?MLGcr zkNRT`ipl8>Z#NFLcq;%X4ZXH?xEFH8U}xuE7F2SW*I|MLJmCk4KYWWUQItDs5f@(q)(P zUh~WeaVU;`ITybuE>_{cMT>*u-^%q8LxvPVE}4>^*47(LY$-(5Ap_!q1H=pGfe7_3 z0N(NV$>v|@0=+0{SACKjiXP@t%ZcgR_?^xs`4FMiU1S zv}WHeITGA`M93Z#YB8u3WN5oK2um!=KY<-mi*6bmWp$K40jZ`{g?O3ndl%1oXs8VY z2iDXdA-#~#CcS0yUckNcb>v3?5DaNzsN2t=ahe}paou8QM1qzB72lDiphs5&3HpF_ zf_`WuD8L4An9i{ejCl#57pYgLW(iCY#2z39YUMB#a1R}h5M<=WYNphzd!0oc(Yh=_LOU?TBkefa6j zK>^5s@TjRkhH#MR&Vj)S{-3LZ^aXs-NI1OqX!<;Rnrp^??WSkkdM|h90?>EBpByLcgd8M9)Qlex6XGZ-dZ{c zNima*s(xOSB-}{rp~1{Uj0AxMn{^)#Wo)E(gamQrzN-eBs~;#vdU1AtX%>rUk0jv7 zOAsh$gy@G{H6cw9n-U;~{JbR#wl^je&7 zzm`nDDWh}DZwyJQ$kvI#j^7AM+H?=)g22Bp_<~9vZ+8!b2+kd>X*%&9WF<%(^1^hn zpBoh6CgAz)20JJmARP~bVuJZc=IKqBjBU= zv%fQv{K8QT5EptN$08RiCI-ao+vV5Re_P>Xczz_fDWzSgaQ+Rp8yCHu!*>{ z(WfDiNrgtk6$Ta$UJpJpv*8RDcH;+ZqmK25FuzC-q~=4s-+J4BmPKQP7$7bD}X(7$GXQV86`+;Z{)YVcGCkCfO`my zV+-Zp%5X%E0@y(N!6j@RPjY}8-G_<5hanEMaXrQ(1w>D*2N^>1i6T6vA0%NpNbp?b zhaf&x*S>{Qk0qIS^$x+A=xCBHJ|7bk>1_VjNci^7BGc=Y~eTZR3xuQSSi) zK?**n>1Z&h|J3q43Pn=}bFHV%-E62aqPJT=~4--#F3vA^F7$o@teYX(FpmF6AM2C(1G#V}@T75GoJdsOFjizotZIKcNp%(WH^ZKYgpf^o9FbCvz4Y-h2ygg9Ff@oa70f_=kY)w~hUxMY?~vDM7P zcyKuSLF@Zp(cx%arlQkHe!dR0({fke2d)2gjZBYUv6MHo-K1lvs85+CEv@?39r=E! z$b)_RC61H*H#U_0J+78#*||L|XJHJFsr}-Ko&2A|OminYd^bkc7!$X`LOQqX1z9zd z*{Mcq_tb6856(UrrfYiEl3N)kSgkMIGn4y>H1jYMChw`=UD}K+P4uEgCyrg}F>DVU zQQI;Po}xJTPWJbVI9{`Vc1zQafrK62Bg`|B>_8noi^phPXSdO75gRdtrWGqO)ch#GN< zRD5=zCcMtbCwGy!0FMy7b#|{xzRtCeVNP6ZzHSslYIm%Ob6k*j6MU^gDIeG>5$!pX zrFn|j96#bL?j8{#Hu6G}9_28q+xH7)o^fRhO5H4RT6 za&Z8O*}a!UZtG3a`=^K3XHFE&s|yhJhre}!CTPsK+s7EUt1XV#=H z)kg}3NGpvRRDH>F!XohxahKg%$(fm%jJL1DD_jlkUdmlc*yl~#y;u59I{0m#&e_Nw zN1y5^!%MN=Jt6Jz@leSD$w}y7&BDCg(~JS{SP4p*&cR6Vi^^v7I+*q>I6n(+&L$s$H#+FOUvBnlCDg7ThPZ zG$I0#ZmmB(_-*!3!u9`vg#rBv+WF^FGO%9LnlYWhoQ2pX+u6v7@J`3_iqW6aO8Zeo z%$R@?I!RYbbfh`52ZhSQsjf*}y<%918^MN{2Dc5SJ?eC94$&+6IJg)aYO1??=M4TJFxjy<23)f&C|+Dv6OU0*syzi25kypPhdTO~g-5ys%0SVl}th6+Vh z1*!ft3%VCPHVJh#sKPQHSGEqnSToxaxbQr5sJr5+&YG2jm6ZqO3XWui^mbb^Y8NdZ z>8>LcC2)TXiW?c2{lplkIO$A&1h8d=6Oje zy}BixfjL*3*&IIbdy74|Ei!y`85e3@|6Bw`-7Nf68=UzQiIl#}Be@(yE*d|{ff5>k z!p6Ul=I*z1!)c zW|31`E`AxD@zJbcCF6Au!St@bKXQ)9(G?c2iiPNPVoKS1ZzjvZ1CwsCPD#ddC@Fjva z1PnxF$D7-I^nq8}$7$aZ_K$XjW`fF5IRd7EkO)F3^I&t+I;fC21^z~v{=TBm(>2%d z^Qa4=3_LSW?YZam;$1&b_YYup&ukxoyfzH|wfb5?MEpKTr zWa;Jujq5VC1h4`usQ{YxovfyCh9)8LAC(QxlX>pt0PY)A)b#J2Cb=mWtQ;#s&tU44 zZu~QG2~@#Zf&|WpNSB4#y_8-n6dPW)i-*V>+&iuEN@>- zzvDZEM$<$Q%Ff6A~W!AkIrWV_RjNs0*Rug#%a<@vzn z8-u}_#LdFZ`3t0Q=oP~q8RXDdB&VRlghm8tgM50d7?v^@eYlVMO zuqN&0>pFxxH@ws*MuL_~{gPr;LCPaQ-xA+B5NtFo?v~-lKpuLA}-I3 z4nhzFA^wI`02ylMq2%nhqyo^2T}k_vWtd4vi;6U;spb zN;fIVXmimhSH0^sRqM|WEZDU52;rTThFcu$6gn~4XH9E+^hxaw#~SABIF)Jce$WP& zI0P*4xu=as($-eq_q*hsC30c^DuK^}?f0sHlc03SZd5hJ^9CJ-NcFUwcjr~utOpN39a{yfHLE$v4N`nsK?XReux0VT67wJH^P~)nLvMjtch&V z-)C?r0H}Y!xDdBghw>ccQySBt9-9?sWLA#zj#ggsGYv@jYwKqS_=~>n+VlcJa4g%- zRGVjdXUwv|P0_vwU~OBfyzFsw(E!i${dk(@oiNOHytN2N-fB-k5m_Nd!c55pqF%c4 zh(9xT*~z59X|4P50ak?$h6ql8M|>XK%<=dy={$>w=GF6J?NQkp;(r}y5MalQf=A!2 z5J_eoEuIj$0lo>@u*GESI3R(M@L4Xea|`l*YS&j(ax5SR)NNh?M$@ia4^^2d8Klaw zq`gv3T4^>|Do(Ntzbb5H);$>9jLl|E3<_fHpF3bNm9Ha#1<{Dykxp+{YH@4{>P)yK za+H5De1z6l{dmhpM#tVwhgVL1I9gLtq49%Qw=dk{Tj$veNa=yRsm#DuH?%rnu(2gj zZt;_V{(fSk+412uLh>Kmy(F0a$)l&nsoXfvO>O(6ffU|^$!?ZotZL!e|N~QQ0~d;6v7c?Kg9|W9c1=><{zsog=pmig3Of)GhA+eJVumU+UdzLx(Q{y4CY`pUi z%#kIut;4jGR$wY+Wa4;%3OILO|2aHL0AMGKgqRZ{XjkF12co#;chmz#fJ+8I{~xB-Y6=~-B3T9Rlu|p78U@sU=;QgivKq9zs`iE1$;m z1)3M_UEhv~K+jER+P^JXSeie(Owyg278kEZ1#hw9gSJ%4%PTi53agMBaWpqUBa^~H zjCd{{X3X{`pzXpGs(YM;-oZc8rkcHI_P1xl^YRgOx5a8(i3j8)Y(_EeJ6rdz7n_dd ze3`yh&W@J&xG@Jy@n4;^BWFsm4jhg~pPPa)O);MVe?oS_8Ob|rqIf9W?CzjEyr=!? z{i#6>*%t#DN&x#byyIr#-wId@OY?v2LMWcI`jr_N_$0|FH08x2x4hK}wjawHo7=8imgn>^YcgLVs@4}n=+lN9OUQBw@F z$v>wtO(Gn*6vz!svnMOn!l3E1rTLK*AQrZUD z5Az8H;N_wBX)Mz(lgkcLSZ`Ns|9I6b$-2q_gbi8CJn@4thi|*Swv=>4B>p)Qn>)w* ziv7FvD~=jl+(ps2mLHv~Hb^9oQ^bWFj}CYL zxz?hHXO01U!0ZZ5A;*)Q)HivaEqN(anWe5nC@ml#*fPi6D1%}MQH{<~bfedy`o)v+ zn1^nKI@v9=#g4OdQt}fw6b)bBW2!WceAG4?ATD7>*6QbJ1UVJ$=163xv^z*YXXRQC z+1Swgbe9vv{|vA=p}1pk_7viDhrqPYy6Hk9La)WYQT^5!hH7BPhGy$t*V=u>-?Bij zViEh}7EZ&xE$ zU<#VILZV$nOS)P40{_H0*H{g+d))=h`lBG;r(hDI$!sI7ED07=Giv7)`wss*cJQ@2Vb(k*MsS%mm>b&Ch5|@sKGs%Dipl(`)CU~UfV_5Ox zaeYqZq+E>TQV|#D#JCSu9rK80`8Z3rD1>F>au0zckMrcR1G9Ul4p>C;j_Szuz=!T7 zRC&vCg%+G2x>jQi@@Ed07fjzsL1s!rOE%!onaA}X>k)<_}?HOml)ao`*T0@o@ELxo1eEc+o&)fAZr_-`Aj~nG2l_Y+yU6OBY=I`2Tnput zY2zpg$g?|zM%#*`oWK+rk1}`O9KB9rO&?gQ+(j<_y&Ic(hmYi5MA#gy|~}??S4{+ z+}@tnvt9At{^WRx^k5A=gq$hl%vf^yAbyr;B_c0B41@gAqjt=HX!!@^zy8H9_l^DQ zD%WO>ad?XhM;Y>r;rj+`LPg81Fb9nI>(kGGXV(fMoZL%Ey)v$dcTwRzU230h6TV_82_;$@tAE~UO4-`4))Qqd_TZO z^Cj4%dgkz#4-YKVwK>TkLr`R5ZCj5~5e8$s`FcY#oYp6~Cgl^f{)j?J&x0@POJ32PplEfuWP_O-67A&_F1t4zav(>`3-Vt%(lf$xGbqF3JiVVAoJIcCJ|@!p~2Xyz;CD6~qKrf-e?F-|fj zydD!>{f3-ah#I5g&i=hl9L-a6Y*7CbZcA6|s~~K6jFE@oUx@2XWBBio@+J*7tR!hy zt*Tx7^hl-}xLjPMl_WJHl!}K1T86uRMY`dUY`MS`%fhx*+OFcZ{~JCFDgLZ*At98A z&Nz7(iU!GsEwi#~`1+bvPTf0xK1&`$Fyl7_3p^Mi_Z7Da zVo`V=a@*j^J5Yv``D)DP7n>^pQ75A%zAQ7mwl+QH&@&w?9#i^2GFEIn#E3*9nen!b zJaLM|U?fUhIwcK;ae;#G)7vo)-Oh0Tkc51qqR`AciyfAL7oTMP=a40xQh!rr9T-J6wV!4Lu#71e0`$Wb7w-WdsIm)ogj*M`Z08arUNq*P=&jF65B!TOm$d&t3tP z-=TIG!ZF1y-R@%|Y+}>tbXx12Nx+446YGI$6_M4j*rqk)uECei&V<*KW$jlnTAli` ze=f~N<1cO+5^L~07ls|iNnRKXMvB?P^udkvIiG)lq5VI_q$|ZEdEEdd-h*`G#p#60 zAp10eSByt8CkC;2(Q0>b7C3n1*PJMVjllMFILWnCP_|N)1|1ZBya72Fr~mnwTClMKlZ5*CS({@gaZqgy&7A!CXC`1L`c?zjK-R@bbV_z$0-@l_4KuMoc?w;;cCu`R2t zP!FdCyivL#X<=*wq~sU=v;utd!m% zyIaF4pi%M(DI12l=GZ_fWVUTlCxrw=*w&$H+{xYY{+r(#SUcS>cc^TOiR)A3cK z<9EzK8JbqAOU3lRei!?uGzeSLcSfG2^#$KgNyGF)rcU_6JEj;z`Fa}pHMn^oic;Rg ztNLOVMH@H5ww$nJlmmY$S|e9rBEIqy%+#z=!P^yYh?;BL^1YJtnAy^;}Y8p?{3Dmcqls< zu!Grguo4s39)w=_paG6;Jg(Z`De7jjjB+x9S2Wk2j_RO7ctfDH@R2ue)Zm4Lk-owS zDsyFLopML|y!(1Cj4jNT-a)>zTb$Ue9hnsz&s>zzrWc8t)!4P=Rf3JST;;mLdz%?X%(K%OIc z)%qn^Y8^Cl`Je;L0`Jgl#8G#LGIMg{j?I}8sAzBFi8=}MDcYb0*5&|h0JrLN+K5+6 zmkcr-DQ2XLeJEKltUZ6o>1;oP$N)+h!@)` z~-C-LP1f+W5AizGOkL z&=>AJvG!IiC}nlF@mHvSc$e{M8s}A)mfzviUApgY@ilR246YjTVV{K=cL2%Ywo)Sr ze97hp<;seFRWA*j}9)wHs1m%AmER|QlQUp_e00LS%{3E;;q?F@RQy!ryl$~ z2CieId8@HQI+RFNxMpV#xo67v5Dhk!Ij32Ec`>9D++HY4I<7!ysoaTjHkdlB)&(U4 z>IDSJ7M>ejFbwwda`6M^90FcBy*pTYGpkpK4Urn}a@&nojS21&va|?U<5&B#)BLE< zChAMnX@_b`v|sv$uFiP6e>a`gS;a?VL*V z^R01(SGR?pX{4LcS+7~_Tp?kx8$I5A;X|#NFOx;lo&C(6Nh5rBHhN9mF%6pVoB^R1 zollF!u+^X3Bjr3XbqH$OZNNG=(sY(nQ&|=uab0!%>h|8rhRnKPN$vxIr=N8wAmVZw zg-`azoo$X59$zi6=vDTExFDm3lzJd^h(u+J7ghposo!c4g#J)crrvpnX}PPF&xwA= z$9^OusFK_X4YLNTkL=3Eiv`N}L-I`d!+%$gq6254fBEMu>xWJ}7Bs8=G-}lPX*1L+ z^Cy&Me5i-=C0}C<(WBIEs}x4cD&mgtiBMyf-Nsr|75q~#VRxF1f}=p zmKju2ozWJF{Me!kBis<#D&{$KGKbX;*K5~SaKms?^tXDd23mQG+vp<|cD=7`+xK*V z*ZWqve}L_AaA5*LBGZsNlqkwY%n!3~F^x$|c3S2oC_(KR&%JN=qr0YA$@#H+Su^R| z;CMSNgva2eV$C2oWQU%{wNBT$pUHq!ZvIZV5m+BDF@>C&4^!RokVMh07$6frzFMN_ zrOkDl_Nv$$3{FCouA}!y-c83k&#tj#_KPSMyt4X*o-d!i*unM(*xa!02R#Gl=`%}h z2}E5E0C21ncbrLj(|dT#og?b}@=}~!ymt%u9LHEj3$m7*m)+!>-wlrNzdf*zNuny1 zof63iVk>N1y=lR+yl9|VM)ASP2>pw{H{6IwIuIio};D(RRXw?zDH{Vz?jR@b@>jt>vlirMn0vtd`siTJR14XyZBnRP=j3ri=Dn zjwN8}Km|v8wx;%0h<^LNv?En0d5r)w_Kr4_KT8A9-sWa&`;usZ4)ExOwPqmZdpT$= zhmWYkc9E|Idh=alK{`oPni&zc!7gj<`Hg>&hBW9Q&lTNS_H!qB>PzlQ@Y%O9ywRx^ z1zvC&2K%lS`A3u~+pRCUe|_FTM&-G!?-Ym_YX+@m zAz)!5$;sQg%f-|uNucbPt|Fp8GH;^q_eb93{>3)|Dw-2KF?6C|!Rp~5Lz*l--X8$M za><8k(-7!$lI7@x^|=m1P-&;RR%5+z0F&ib1MHQcccwjqKFCcU1VOUFrb!5QRZU<0 zls~|8R;R@WA$2idavz8D@XVC@l~~me@oQPM7~bl#141QOj`{!isbQiC&?D=_3tue? z%&bp{thV@{=-cMN`YEB4Uh+;I-I(gKwC$8QecL(u{?UPDu^{X&db@AQB)Lg``e5Ma z5vvFf-QXJ-#_xOS0-A;|(q@+-UoKzE2JQm!kXZS#PQ=em$5v$Pm_Apb#Z}?t4pwjn ztB{mhEfX)Sq~O%Uh+t6TG#wl)|b2i^~SI?KfFUDil&xo&|Wv2J2Ae zOy>Z^zNmp2`@nSSHJI6$hBsO}Jgk4w(seqZ<>@n1YakOaK)macAngu>Bw7GS_t{@E zD#mT$#9|T%k%4Ri($`JU9px}5o8*aI5n^FLD#v$#&(mE?{Ey;&o@YGF&5_l{-=@ji z&9H4>i10mYC(k1Al|?0*|Bl;hRx1QM=Q&`)0m08-#4@DNqPA&oeK-H*HOnb=2YHb& z8>xhubV0Yw!4!wF;bD3@iNr6fQXG?359xo;v-Ucg`92D_wFZvKLSgHcCvPPB`AoY4S;&f~8&2=BD_P7d`%M6~#5O_vYDN<4 zfb`RM8Y)+@*03Znue5*o&}S)If(3UD3}YLZNwkca8J2@dbQ)i?XkSbD>B)1K-62Ho z+cRSuXn2Q6$DLGg?@p*)n6U1zGT%EbE#EEU69W}J)$asZhiTlr&hdt$J1sa>&Q2S+ zZvU$Ox$Sv7XtEFQS7F~y-5UlHNgxq1u#&X6h0>Sa8*IcMcZj(k@3Q;@Ont9h?gs>{ zQK{^ED!hIZ&2O}lzy8s@gnSX!q14Ye`HAt&DgnZmq&M=yiDKv$dY)W(bqw`*#jRYJ zy^q*Wg(4LMY;ET>YXATy`ZQ9Pk~smVg*of)75SK>2EtJeuS29HH`K!$bRwQ*Wl|}r zb|`QzyV&kt>Qg{;y0U15b$uB{)ql5Q-5;;c( zydx8H?bVPM8a=yY_t6jISm80xesph$XP?Z^{#*3uJ!knZFVN34`xn~nrE@VYy{=k? z+Sdq&+Q%|PhgCL(R|>w~H|<10!b9~SMFI5XZk$ILh@IK&0r|qC|4Z?mU@O1Y6i@31 zbIMuS-qa#J9O}e0qWpr2nJ)ts}r3eDCh^Blyn~+ ztf+PwZGB-h`m1+go}Z4a7#T&5>O?o3BRl0lCgS=+QNiy?1~b&)J~PU0vZ4{W^o~&; z`9*d30L3gHC)o#A&p3!ExX756?8Q_+sEs^u` zwGQq{{%B!KKAb(ro9?%7ojA?PIk%)`@0Bv?Z#)404Tr^B!|bd&rBSoXP-wKDuhvG} zr08r}_#c%i7)FZIN0D0DpR*9>de|pH#gfcRWX{sD@uSft(ic(wDF_6=7eG{pMdbnC zfQ`l-f-tpFzzoEihXJenv1OPl`7!|CdVynvC_D^tsg>80cF#d{VJ-&hY-AE-P@A## zSO4ozqq0k7O2e3atO~qy7jQT1j2$lC7R{yFoE!(bdTBMy4Qt2W!Y~LkDh1TxGuf7tF>^~A3YGV|q>8mn3FNj=| zj4oM0(&I-N=^>OJ$?+_#cq^paSl4fXuQlDCr(k_@_}UVSqb; z*!All=H7i_Z6$E_HoNt+hSJ#eNz++V9OhaTALI_Sgi^Pc2`5+mHg->g45VLf!tN2U zfMZTYYgCk=`V?5WBd+{Oo?(VCi$W2%)KMx|j@yS>*+0h}#rvZy`2^086Q1nPIq3qI zwSRzakNBjowew%j=^U!ri~Ub8BYNE)>7l>@0|8Zj*;EYCG9cTwM+9<^n}6vy>Si3@ z7(=FhXpH(gHU2h-0B@-8s|u|qQI;E4&^I>l1&^R`C~etiY;=ibhkx+&io)^CGlH++ zd*OEvUt06PPKjKc{J_b+z=gT-k?D7ZJ!#cgDRA-yK*@*?>V&ww=$uX9A+bqQKXW-y zp+7?1)B#p5`tv?i5OJoHaQyc-bh*URjgd$GEgihoqgNW-JTB=vFh3zQX%;vUu@vai z@Q{zCLja`gx8&R4clu6M)x~80G(jE0rc+9i5-AY1jfAd)@RMTr=Iw%TY#=9 zlWqQNl)qE!;C`WX!A*OaNTXY|#TV8jgL=yvx?h%FEfl=QC)ldN1i3~8o8|ow*LYWI zhUs2l?tCkCFNv3uywwstnFT8^I;=?7v24le7FeNk4kRk5O+Hv3u6X+vOWaQb{?zNI zw{Gnirxf7(|GOE@90I0AZvF=tz_f;EIB>v)SNtIr&suNYDTw@QiBLVwgP#f5cX%Fl zX*`wFsrUCcmL2UpJ*&Nha!J$6R;$2XQQkDmURu-MTfT`x!=SjN{TvzpJ$^BWR9`n) zd``nOTeQ-bUGBm!;`jmZyW#&z+*ZCPB^^& z$l$wNsRT*4ork0G*U8D(Sn@bm@x;x$UP2kd;C55&2fv#N?T)m#`wmzgms4o5_g47& zg#qaPV>;1d0GUzvkQ8PfF%7sV62Mg>PC=N%EHxD@GsUEF5G$zF<6zmX?Z+yxiv;ke%uNhfu|hGG){3Q^IAiJv={mpNE5pxc~Jm+nk}2VlB+Sdn_G{i zYv}b7_v*HY!m`i&VD&ETAopEo)`pX6OBK?wwt32D>^EjI`>-CSpqG|5c#~L9M_ce0 zE$ohzcvU_f8xOH^Gx}7u=I}M#TCT|IA|r5Y!-{pkuwr;aKf|J~t}ss!^*E+*@JTbl zJc!#F2>^{+9Pp6nk!F`5omq@qE-#+9bpEW)nuXcBwZc!53v`3S-f6L8c2@qD#$JY} z^I@Qf>;>5U3oJHmO@sak`NMrXQoIzX&O!BZkw=YRg(x`#zMSFg?grO=)!g&4^yUaP zIty27SL%bq`R`Jy)<$(0Z?e>OEZEICBA5B^+UwoSHONipm(I7M^RTOlO&nB4F1HN^ z$7>j^)lG)hA4{Lenj7Yh_;!TZO*ZY~*%g_#gx$1XWoCnOh43JyoN^ym(y@{bv~Htf zYGsBX(rr1t4wKi}d^DPb6f(sr|1_%pso!Q;(uL$JNx&&5CaL9hmZAw_i0EZyn6K-# zP@HshSQ3z(lG#iztf*|S`+QQb#AW#NB0RQ9@AIo!YDD<3O-(@dIk4Kmb zj5Kh+oK%(G;8l^KE0-SWR6*0)9S)2%x=%l*xjMKS`Lq_2LY;+BkFnTy@Pm>b=PYbt z-FQZipTwQUv})5%5H}zovMp>Oi{cod;YvPf&v;ENgWQJ>lR##wZ67bzX zMldoNefRmzzK75NZb4?+KH#0y{2ZhpS8sDD%B=YJc$M(b+c+6$oGBHiVh6Ujdd?G> zgns`=?%VBWB2CV7Q~gx8Dm3H6&E2X^yi7`LVes1&Ub!rsX!Vke?(CT&w=$*Hk1>cYsoZe3uOdF7+6M?`Dz0`sA}r!Z zgWR{6?5@RjqIr##E%6rY2=g$GAYfkFA4>Ph$vBi$_5}KA@ik6^u@e@8{@P`bNgtbD625EU~)wo$$lH5XQyBs35fuxL?y=0LI~U4 zwR_uM9G7`&b$q6?nYnaKF;_HS9VWx^1thIyEwe4rcy26Xruc!3I22W4@-;Tps5TuZ zT2eJSaHa}t{OQ1>olBMGzC}(~QCc6)fjk;pQhPDc9g8#>{<>vc%^ow_JYdG~B@YON5$ zbfpbV&SL0LDU^0siCKle%wEtYK%YWZy*WXk7ptZKu;0lF;|ADL#|vwR=i8x^MQ*wb z@es!fe(NSAa=>gU5nP#h5M-oRF>qF|WK?R@2)=YLb*&CQd6KrFSJ>cwfoeIaA`x@Y zUH3S9HF~&~Y4hLock=tR`M0*aMCq-36+Ux1B-i$-!RB8f_sLKj>SPE!WVZ?npK^*! zsGaaUJDwHl;8=ytj(H%~n`UF7SLQZY+F+rdf*hqXy#_E1N|JCr80olJ;#D$OQM+u6 zj%jY!FLGOc9EqZkdov6g;C<6QF^7Cetm?VGsFx?d7OMLa^x50G>9}{XQoOk{>8Vm# z8K|Ie!|YgmbdG1bzVTQP5Ut9h=MO%LTgiFB+MlOQt%5E5EpGp@Ezk(BuTtS;IPS8$; zy3uQ`1lHiQ4AFKaoZ@eK8OkVKMdf$!tdXdQfmoU(?9@lrcd_tTC=;5OJfA^b&KXo~ zL%nX7xFV=C_*OM@On0DWdM}6?xUY4tQRJM8FBEF$32HO_%r&Rw5oMMP=wt1on^?9$ zR$9*s^bRmh+(LzKA(TTCkE{pcTXp@}!NjYBX#$_e&K{8SGJEuzn>MS*esu@$px}GS zk$OdLwe(2mpJrPoqnT7A8*o@s-!mqw?VJdJw&nBes@=nS+Qtnh&ht1GkBkgZ>Y_8K zGzO2!@W?>zPjmbgNYAh-p|3yZME)g2@tMKtZ|Ow2+5RLLp+1`Qv@c(N89=bgZQo3u zlVetnFH`b5Hvjr4MX2D$mWqO0!zq5%wLeXjl;kRJGQ@~=g(zeeQ+upssJ1iztqXW9 zi!(K@9%Ga>u$nb{a{E+!G`adAtglzLqQ9h9LxjJ3MWWg*Hq^th;=V3D%fu~_7QyWq zxKk|^-Ei|C5H_&b_F|Qvumgd$p?~$-sRcls7PO{Yx z)FMyYI4*5CquNH!kDaCNKJ|}?g|?`DIso3%1ntl)gp;UqGn#@|LNC^4;f;$tvi>ru zp?&Aq(Zfh2<|3u2&jV94#8VdL2ZniaY|2)RE_2HIvc>;78dz9Z+^=IY(9(#(+=?+j zu5=-dno36Y>MRN#H%2eLBQ6w|DOvL~fXg1j`%B8+g6Mir$^;7O=DbR{`&xM-M> z_5heFFQb?oOKl!4>bRqHCdKz%oq{T}>K;gm7B+kquNJJ`e>%=9TZK2SWR5V1+%*ll zn#*9kjd+bE_gu(x8Xg%8yo0LfHW^%B+KSq}$SzMInb2s%-BwP_j9}^a+{>ZR2v$9G zEWv}y#(G_TO!hNTl{Ra;x%?QbCR))p8ACc&Yh^letg^z5O5Ew;%w#nO2fFG5A99fk zZDirMW!2R&q>X;AV>t;&?xEbx=J_!2>b`{N?9&KF-gqD;+$gj<-k>_MU?ZM467$!w z>I&z~$!PAChrk5Exmnxk#>MgC&n>fgZbggZRYhYXHt;s2k?ITj(-X}Lai{v$W}K=<;zpH(Eo z*c!|&<-ecS1VM$pD?@+D*vZLi3HO#%jdV+fFC5hV1C~oD)x46EMv)aG2Ys1Yu?szM zuE!7yzA%D zEM>^-8Y9%kWC-bi3(>2fdR1*?oi}RqtDvqY@+#v+N*w4V7{?W8YS-mq z+i!-l3{C!&_uC37&|=t4M7fz>s>;JPkbXab`P|kTL zyOQ&RoNnPadq0nNtLYrRK7FBDf2IkX&w<+soP~uH(i^9R)-nGp`-T~US!ThD&Yz|9 ztF2`-jLbsIy~B1QWY3MSKEP8W4+sY6=xiM3z_VYp>SGm zucqI~tNu)belDVR>2`p`4sU;}L5oHA?Ku=6r$w+;mmrv4j|D4*wn>TGlIv-lV-)|HcW{@p7vJ6jjc%) z%Q@J@K3M=wI~NX_?{2gQDz|+tzevFm8!q3z=sew!D9S)syG@C^}(ZjqTQ>2c1l=Y{skqjQc+ zRl(#nr-gYqZHp9dn-Ai$mjgnc0SClk0Z@%w8O!?|GRX5=eTKBY!dtbP@?6)U+zVG% zPUpHUnM!2-#B*BTc*H4LmBQQn_@1_VH97=R? zc6zkYKaU4G;o(kds>I4lS7meeHD@`otf-4CD07z{hrw#4C_OlJDOA52K7cX_UdC=N zLzMi}6z!g3{QmR9HE^e`cYc#@Gm2Z~#VA7VVDbFlF)pBcr2k3Q=+6mCcxmN+uv z1AAlyUyHXe*^6k`zU*O4k2lSzAa}T)nY!767yqaVT&8TW7rnij?5Y1Z)-fY`FReQc z{H(7z$gu`uxZD=!;()<)9%?WbkZJeFpM{SrLl}YZ3pQ&g8~TYKF+Wk>!xrS`qo5w7 zdAQjk148M@?M8}#vz;QaqBT{Uaqd@X%WoVj4ZVZ044;R_yXj=<~XWra)eSXBK2 z%xIw!eV868o4cN{s0mmuS{tIBsoZl|1t)E3CIv+gV8~*a0lxJ0jteNf68+jmW^fgx_6loSn&LnPN zq&I`Vw#<@D@aO&ke+H!eck;B}FWm9wzdc3V)8ee~xut!K7G4f1;0#Dg`JkT8J12a?FU4VI)ISX6Dt7cDGL(2 zt)$O;`~-zRU_f|1$7h|HXQX9~oCf&(_K;3GQfldBA~R&meNfpeds4%1WHz{cnz|J4 zFc?Aai+_Azz}-&n>vd}h=o(miLoK!k#S*G-l79s zs=Nha!f!7T3v$_-#mUVMMRg^67yU4`Pn{8s7dvM@*(^#gRjsLNrB zAS^1fch4m`Tl(uq*3GK%I&pAj@+EE6jBH9-+Zihfb~t#C9Tx7Kc#4yQ51L1lg_8k z1S#!sFLwKx(@A%IEi`g;f<#0Tt=^|#waOt%f7iqNQ@DF=XZFmMyujgM9DIphWgYX< zh}thpaW+nsE+2br;FG1QFA^fdk*EjaF=g({ReRMb&O)bfOeD&*&$~N|!$^0T)Vi!z z7c?nK$M^2@;h-wAV`1!p{NlN6F~_Gm_k$Opjc1? zC?XaTKt&Whp#%sWTTsvdO0k0FiajAT5flU!8-^lcL%LY+JNJ9vdTYJ&2Wyda&MvcO z&&-}Z^PBxOFA}%2m+rFZ&6%;&sEK4ZdzM*c6!D1{&DY(K8QGqW{9P%!AFR`CwRXsS zq|=X~(z(k1mdE*yws~}~7@(6a;#b-%xcqu^ilskznUNgU^|K^uShs7+*RKblw^SSc zuqR6UjJAJGuPWzp|?{<$m?w)GC)o8gWJnMR*Mo`=sxH~VF zfEDd+dboAlx8xOSnf3f4zWVkkf3a;zOnQXAYL)SZI;-A82U#(h#gN#tcaGI-mD^5O8keIu7B4{EqGpqopvz+H8u3=JZ8>_b9(+yD3(E!HhBKb+qf+`c_JlGXl8iU8G7+8YJk(X82ubAx#a z;>$}z8b_U+VobZ4e|Rnt`B77>sM(6VEflOFP5Pbp9Hk46s^k9!w1eqrRRbtZU0$5f zPb+y-?MgH5UB#sy^>~VK;u4=s<9RDG8~ihhR$-p2RSe4k4}1KhnOA^gjwt+z!)9~R zMFpR1lYY`v`o7WBHg?U3F_H3Tv341aAbY><^(&GUA<(!)aK=}2)W0j-g^*ElFHDiKOLqgjg&f- zl}#=yd9|o)CY$6=W)+uLc6QtCC~k5cIoh};=+dXkSIjHW@Z={W$D9^FJyPbIfcCBy zr;HAVlBeU7t6eD4%VEKdO-6Z6x9c0@NX;X&7iNziqJOKOV_Ml)XMSNBFLtsx#rFM@ zB%dcWY3`V<=z@*wUCdW4x~3P|zOe6$ONBTTG;U1({o{n?tpA}Oj7(ZF*1@%$#eFkB z&nqHK4FMBk>p=9xU1;!Ae%h4q^Hgx=fwU#7UAyirf9y9fx@$qBkwfVVy#sF@s%%er zhc@<8o1m(<+F>{BOm;hDd;GNXy(zg>v99YG(@9#|)txZtFx!<8qu+$*tD5mx9P)Z@w4c-s@H zuOjcyojJ|VOF)EQb>rsv-ORNYEBc(`oGJ!bQwK`gV+^EGAHVbbR(l8CL4^GGvJUFb zjyNsX1^wOCgDEas-a9{24j8siT#!T_c-C=v%aQYpC41@DWAwx(BT2(V%N=(ah%@w8 zXE|ayj3aPE-P4>kTYAEs25Y97CZV@Vn=OsoM`9ZsLbNvI{`9+*SrYuH#Kk1wN&RqG zV?;uKKDn{Uv}D0|?MMW^A>W!ccsiouzx=4t`}RNn1)e9?U<}^yoLME9<)~>hW&(vLHr}k{HT0@3U$D z`|&L^X!wrryzb!`VnDa`8B);bmR%jP;^NAAlb?2*ESD5N9eiB+;nOWi{Jeowa@=hj zx0)Wus}ns!bLqIt?Ezw0NT%RlSW#JsmG&8qh9?fuR{OZ)b3{xm%0y6}cc z%;J0NOFF7QAEOGTlO`tXqzG4xBvKO(+;&&czt@RsW zOW8Ae-@K{zj*+9IlYRTChaOk7$JyVXPPQD38WbZ?gcW&`guKYhNuPe@Hwu3yYX-D< z`+VFPlB<16;l0LvbL2exEQyV8RGr=DIeFz0N_plfY1wt_;^K6l_vv$d=LgmFOi9a6 zJwe=faqd*=i`%=Mic39Bdm0Y~DRP7A-&eoIOUSRxTgE|ZYJ?@TJa+Kcvq*CRK>|iCBXr?JwqR@GKx!*m-`# ze({c)&VF8hnkKD$Msjz06*=Fx%1)zXw^mh<-hu0^j7I;C3m;wQ2ZS(w^x62~rCohf z?ZeMHg2zL>n-=L5PkPo~xKNGwqaJ6pRz}i$ge}8Pjt`5)tj4|r6;0X^_zd#>4AYW; zpVg|oq;}uvN#m+}W>TGd4n49389e325qV1Vv*(>}m$AUmu0o((;!s&rtgtV3A-&FEliyJLp^-r&g)JCB+hgOlXZ z$k7!EkzLh;#bssfeRemhA0i+WaXY(f23+cfC;*EF8&;fL5&a`?>9B*@vfLKdkDLw3 zBOflBTIQXKo}6{8Dy0;6oS1Av+1>bYk7dV6EvBX4Ht&$!j_`|&+&FP7Gje;+dtd8w)w?sxX37S_W5~0T zQ<;dXv7{-cW?;X;#YU31?{!rBs}13oT_c}&C(IP*?sqCKU+j9M@nFDL-Oj)D9OF>j=o=GVGhF|G8MnN0 z`^m4aQ!(6xVyU!o?*Y!TNe4&Bsw?=MyDy3rLbq>?@H(eGy$r8eVZ}YQxFds=|b&VJ#m5#{o#xZkR zLMmz!@6X?ujB_UktZpo8e8-hok8duIZ*T0|OK!LK6}nbcjqdmMTQa`+b&$P3p}crb zUUB}8z9T0S$}e0R%tP1N{Ad^6I&ku4l`9{KK3z@9JwF9amOZ}YdA+?|IAD0B=bu#G zspi4iQ#Xxs+8xmTw)zQkRaJ@p!Z$p`sN1kk>~1mgNhhaE`wDkSJ=8*Ds!JsR!5YE*xf>2 z9%gCXf*+8+nOb(NtH?W__s=Uy_DOn;KVXD96kXz;a0C&3s zA5R`2#d#0!2SL?UC$LkM3q`(`zWFEr6EbVM+V`5K25H4nqisQT)h=x^A2)+tH+^+K_-mvDyEW zoNAb7kwH4dbvz$ukkyja^Q(M#L5sh4Vq^I}z4MC`%H`vF52WYovv&B|Uw=H${DhO; z<{ftv21)IZu`KZPBrxN^zive5Sv)5NsXr*eFt<`I3nztYBB zvRRTcpE%?vts~4zjNOLAe@f5XM(Kk=66+{?djiHSBRWz--@YN{;}*2x$-TY^l3+CR zK-Tp?-d)wJQEaVI$=gc>l2+W8Drnkc*Gx$m>BMx7-Q=0C zkK(Ky67GKae&!Qt%j9XlSAt)MXXYR7JiI;YCgztT{F94!Q@4UPp17_oHKGjv_@=kx zmq*vUycqltKG{Ox38V5YcNw4fme9The2B=L@qDo2M|BZO-QxMdngiR$aui8={DUd6 z2OAiBG}Umg=;Z^J=^me@KXO{Wvhv&S)%0xq=KM(dwP%lG(%%Yl41$K|5kb0^FnESk z|0emG-DqyhH|nAG2R)N^#=)AIgDmpo+894({p?f3Tq=pOvoj(KTbAd)GPNw_IZDE< z4|&=#W1_?NW(}Ke4wD8H+Br=H*dW1W9L(It*Z-gvL(~0#(n2gYhqLE;?w|9nHavy3 zweBaBv+hpSOOk25E|MhZp*l;Ytq>9~uD)hKp8NiUVfG!-HJi#GPY(>oFtVmgX&(Et zT5O+7EKAn$RQTP+w|li7viI-})ghm)FH7I{?sCX{wNdIOxVlvvdUH=KBzo`ztRjD@ z-Mj^}optdBnB)=g8_=XHB|WDZ;Gr=ufHY99z2Akn6`W* z>Cx#lc=h$$e{-^cZsLkQT09JgYq)C9*E^Wj+CQnKyTsqsd)YuVy0-S%8R@ivich@T zEc}j&br{CL_bPtE7xPb5`&^Y_1U5GPFItLnbek6q{#BNWAXV#Kzuzj+oKQ4MeeWS> zjw}OSeeZ^nvXen_xovEp@`q#5F}^rS2GxHZ`~g|L z9$QiPD@Oj8BWDlz3FHaCpbK?KTrRH7_X8<)Hry1Pt|-p}C@T^FL?(2^852HomGv(u zYFpQ9i;0=c{5X?h<$m&BWowMzgcU=rF5MiwfAo43$VHl`LNDEI?j_lr&7`0%wcMSX zO-X}?5InN-!`J?#*&I2}A7k0B0xkn8{PeVrsqa2LXa4@^4~|DQa?>9OE(%Y4N0DZA zCSq$o|IXYqeig~4n|We3WHGz{hE66*v$#$t)0c)+HT4LBS}A zUh%vOfA~?&uEk@Y_|%`9+xCzL{{?CIha!2{Z1IM;8{>q)~r>KQ>OeuKyG2o zw`bpb(E8xgnTM~Z4$n8vJJWhLe+*k@!nf1=;AX#O?zQ^MgT`*Y1_Mx##Q+*Kn1g%{ zA>GT;=xjDt1p_3~4Z@OI zfTyik4woE@0We^dDF8~H!1pAour$C1JYO)F#$YztPqa}uzBF`}6)F#V1PA@M?Emw0 z0)hoHavq3nB|w@4pvKb8L>i@8YFjZ<2^Oi?PQ3lK4jKxFjqqTkps=(Cgl=@&e}g9a zrxQ?kvJA$SxBuV92oQ=8+d+Vlkq{s`4q|YqDgq7Q7;ymbVbG%?sY_64x$cRf}p;l_er4!MbtU}~7Z<_)+m4_==+u(SK6%5-v7-^w(|rv?ooW@ z(SEnh>}F?O4J)N`eeg;DC30Y)P95-q0ETGMMr&X@Bc?CxZQ=m?>q{yA$9!52Ju2x8 zIEDp$48b3Q+#6U>2WTz~#&D=x&k(q7G*pLkAW};ulgMC$2|%y)pr{IHcVJQ~1(`5S zK-2*QLTPE@u_CEeC|&IZWa&eN)3HiZYycp3r!inOLAH{iz}f>gRrWCJx_|L^buuI1 zoxQMYdSk~~N|+BPD?AgKBSqy+A9sfa5WYhez>}fEVC~u&JQS9pq(C{qpJRA}5*2GE z3{3ftM4d|jTY>Nopt(`vHL(gUx-gxLTE$8gAh7XWlVJw?M=s7?<=Q$-S}(4B@_MUW%0 zu&o4CiWQR!w!qaWb%6E-0T^1!F#(L##~srGm_;*KDDaV0r^HHyBXe?yG_iA?yFFLS z-vedUjb*^<&Sj6VRSMW#0TgZ6f=5=V@lYUx0TE+??_3zB@g&s2Hb$o>IWW(8X}dqK z)QaKGkYgde4l=;o7mC3@zh>EhFxCJ%5qXORu_KQ5Trl@Z(v40wcNJ|=;@qk2EYw<* z>Ol*@+3U7omx%0B0EL~(3IK!*k(;hNc^oEiGZDE;5Uj$QW23gtSt$ypqgtKI^Q<^c zZUHGa5vf2|N#+t6=#_*cZWIg$1S&}LUicwmQziy=ND`3};sP2#6TlH4PiT`vrZdUl zKp=WekQcS=waxn$R$MM+5KCW@LqN6C{DJcel5f2iK#B-)@hpop2E+zvF;>6ZXP&YC z5jHmVjeA@{tybsQmmvQh&c=x*$iZ~mkqR% z_z&dFT?H&SHnp1c@g`aTpeR-o(5qOT0J8L8G;o|@j#b7<7|msfR?(O+XaI=7{O{>R zan^=J$q7K=Lk(ylivXRAu(}WHDuH}e4z9J*3l%g>EC=&@KJ2@M`2AERRi2#nVtcn=vjxw^D&F6hH(RC7~ZCPNq_jau#v| zp`SqP3?;J>39s&)l1?jcAa)UK`n9lVbu8+!6$2)7)R7@Uj1)Yw zN5Y%O-2;NU1i+0L?n#FHsBuJk0gOm~st~qf0)pN{@Kj7QPjgm2Wd+c&x)Q1^g@9-T z2tddZlp`TIh7eB^6ymT7B^@QUx4<&^B2a*(UPXZ*9>Fe!@Dv~g8QSPD8*sD#Qy8&6 zHnnWzKAo1Dhch>3ba?@Qrw`;~yM5`MDYVqb44wiCZ)S0G?60LevyodmHkk_p0iKRY zA@UwF01leWqoFliWJTTxqxi!hoXyb9AY&fdaA0ghJBZMr>1GfyVk}R4<>J|cXh0rF zWppX%GC)BXlATFZt=v|49AHGC3kKb6GDpkVUqqmFK;}AFUOqUILD7}dk05A2EqoxN z(6qo(TH1sGf`VREH?INQn}CN|(8@WcvPYJoNZ<^-b% zpdb@8NGK_ffpQfvZ^bsSt$?rA3x9R+g38tcTtKHiqBr=`xC$B#rpk!Uh(#iNM32P+ zbSlBF0r{PUL119f_T?DR55=b+*PnP1X3=P+;$mq8sKOIsitURq&;$}XszFz=?fG14 znUZ)_!K!;g>qM~zU^4Jy33F&%>0W*PU7WnRjeahw6fzqIyj2uV25TUL=8ttA(_(aa zAWDUtqE1s1AvwUx{Xl@o{y%QRP+%=0xm*3l(cQI2Oir)~G&x{sLgSfMdTeJPCziFb zFi*&GZ9+rVe;SCXpks)Dp+;A!F+_43-4v)fV0Qsbi<8q-t7s4+$i&W|nHtf?Kqeo! zih#)EI5zS#?HVEq8wi#XN}vW1rC}{m1QC|u2C(CRrbGma1_ky2D4cVlrWGgOYMemF zU=X1NpiCOFRqFigWds`;pw>yKXc3E+i?z(L(QB}R0WqNAKmHV{A^b&ez}S$y>^rf} zXf0M(2bn8jK$?uPhm0b`C~R<6WMaow(u65g(8elAAx5~<(HyJ)q=642|LBP6tjeN7 zSfT<8QY0*0E+BQ-%X4V7C!oQLz65VCdI-U7R{qFdg0YMShDb!D;gGyJ0wTd9xb_Vz zX@y?S=3aoIAfp?xBsrqhh($iYE|aQgr`m}DLqt>XSPVQ(P9(BDtoT?6Y9yjovQm^# zm>3(U3BvLjK^=s+6f7wP>)!z^fqgBC!ovaKYJke4*}=e0ldjaEFo1mn48^q|I*eGs zbylskl4X&3Jm&^EJqQMR$f-69K)x)RFO9#NT!xqu#N07<3&`C|8fDOfEVe8~j zYC4wt5JY6sRe6ZbO1FDR4SPstdqTjI+K>sj3ahYohAxgUmO_xLF~CD|_hYa`lQ4*; zjp4`zo^%v)NNA*TQ|x&<&atblCb?FY-k>&z41jYtn?k}f`n@Qs7Ay}%2=rhq(X!I> zWpP%~#{96BNMH}aIY0v{(x3=hv3gv45rLk%61oNhs1b1RRvdSR9-5}-VP%P;sPGI< zD^}D@VfrF9PMs}LL_`R1I~W!y#(=s_!7xE#=B|Hohx;0n%KQIwNKCs`ea(Om5=%-vp73euglO|bvyN;|dK)Iuu!-R1 zMY?4~JAK#B|D>MQmpLV*^Im&)YpJF^W-QXSo}>N4ym)-;{P#=d+5d{mdFj-P2hr2XoXTc4IZ$@#zm zLcGJ>9sdHqRolBJxn_xDw;)pA%>{0=U0;?Nk2emfBo_lH74M(w9N4;ZDTRMV>^l}J zvD)eowIxcLY~&)b((->R>bz^d&+DY_{DHs8RWx@3WYKmWD~S0cN6p^o-?j#M)~?bp zLzdc{1TX-Fpw|8WOU?G@GXj`l1)Tp`XoMM7Kqb8>^)y=Q2G4fWs*CpXzHFqSZ0*s9 z6*jswual|W*eT35TZ7tEcl8nVf&(0k+q??q7%JHqIw)M7ZfiW0^%nsEUQ&o9_^g%_ zq7pdbTWP#6^?vQ=NUK}Xgn8;W)booB+T4|ox(^wka$G~r?%EKU{fR7eL=wgC@lVk- z<`d?VT94TtJ*b3z^7yEprr!MFjkU?Y*z*_fy}UAZ^-+hm6Jv7CJ7>s0IHOI+TFP+u z?(y%L7bIDSX+gZUsWsvj6*uf>eB{;VI)vnwFJbZS#}>rDdueg>+`D7!Er^S+KJe(d z1$)={C0bwk7YQ?aw&lm-3d8@1^gmTau}ED}giQZSClo|7XL4u(VM4}#WktsnUH#+i z|5sHM+JuC%_3qOQ+ms48?al=Z0pw4!k8)8R;)SPC9HN|gA`ajGpmN4QYLQoU0yJdS zCQr^#8D^kbzTt*yj&Z$05X|k^@ndf1lZNXV2id`mS1vsVIxpke7}_`Rr#3NR{LxC^ z7o9-=#=ug<2a}>UA*{SAMX{7#@Z8`!!v{*H~JVq4`s zxyBd%_IiTT2GvD!YJzD% z8(o?j58<|&KAj!Sx$3D#?uKKFXaZBIGmxE5E#ptMF@`{(bJT1bu+kcGq< zxB&enMx-`8S-f>yD@s5x=}d!U?iT`D?*RkzsiI0>7Yi7+W(-tFw@!eNvL~@B?Ft^N zEN|-%UUg^lLZjTO3!>UH%K7jJM;4Qm?|y8^L_jZ+^r3H22t!Jg61(gSVue?)Q}+Nj~4GUAOtFJ$&T~ z?y{~H>VmZ~o;!1CY*?oX^G@A)%v5gzSBRv-TG<>#KiE4=X~lF@ECej~?cWWWq#D{p zA%C$SSz)jbb^g_}-7$aF_v=mf9fTO=&sL?N_LWQ3AFy#fWzz;A@WAN~8RbJz4nvt~ zUb>o|z5LKxST`VQT8TaJ310E>0|+Z4vitSz>Ef*f0+%f>zxpyjpe9JSxK^(S z=ZT>P6R}PHZK(=YTiI336)g(>!q#{`2WgB&+?NLr7 z-gSKPD|tPp3-TuU8@hhmk5(Zk<5arfm#c%syrZFpb{lBVKn#F0>{nJjY*& zov<3;`->{%py=!=^t<;L4E6yNoSBsm(D_d=`fWbREvh&6_=v8byQR|+SbyePev+;V z4jaA1h6+gs!jGzNXTy`p?~2RWIvM2bUe76%$Kj8**nRewDLpbL{Sklu-^XY564Z7- z+&T3Ojg%qx`_a}X`mkcuoHLZ%PzZ9uGSiy{Dvv$o~wB@{5 z_S^dDu~$maeA#hLo;y%K(&6cx>g54)LR&VfbKCwo%djgp z7v&4zojgc$Jh$=iDBf*JA)H3H418Vv6Xy|V&0lilPG*B4F$}wnTNy~pU~>)@={LJkn^zhirub@T7-cWAk6-Z7FWY}_fY42gu8Bugnq>HVrX07qmdG1rTgeJVm--Is ziZnho$)E#wOOw5R{%nRz(cQx?){f|jwrrdrJrV+`9gwe9ircUy;zq|cqxS=m25Bzbi zjG$Qvzc{Si0u^|>g{crd@FMvA~=-b1Wwla|Q#}`snaL z+4s}(IZB9lJ;#BDXWm|eT1>cpr;vDfe%iTNw(I)^Azb)# zVM}l?-bsk>nPsM*AWb32|HbWTOmu+QQz`W7OJtCX5tD@MZPSQa<@v{&Bqi@qa;J zLqU+7lN*h`abkC8z$c)Y-nQme_7}Kq%g6?EG-26+)HmvDj<9irY)_iAb|9drcM*EM#Hsd*O})U zcJkeG7nULh^?0Ik0{n40`dCJ3d04~odf)eOow4Wq$eO}G8S<{O_WHrg~|kojA!X*78i zY`A!OXvAxHEjs@bZvOV)d*-)`4kt~#Tqn9#dhj^DWR{PXsI3@+HdhRO>dwqM{W@`t zeEpxduVM7`qOvlY4*z2k6F42y8I_KD;|_1hO6%nBX)5)%%G9YP9o@{tT>d75xGmi{ z`8YTZENNgC*7u6Rmo?3VOflwAv6ZNJ05|GPoaqmLdQvBvHVAn#pM zVcGoSp@8X5YqDbWIA9qayYP!w8XW3v<~)VYg#QgaZ82D)#S5yAl`qehpO|kgTJM$m z08eN`UvMJ6fqQmUcV1c??;?`i-FIBu7pJ*AC5LVOspUHR06n(8*LwR4HMb?^iDzJ< zN2Xw}&bNYoo(G1wkFsiDTPt-}wTfe+yIhN>ff>edvY~r8lTe5?lV;vf7|{bD|I}`n z?z4Cz;fSiG=>#UViFJ}p=?puS&cYjzeu8a{=5Lb33<}o(>>*~%n-jR?n8q=(gqM13 zD>D||%2}_qYDn8C;}yYm{?q8F-dQyxG3{q|iy74?)<7wB#n~E9Q95U}Jn*Y1t^0Tb zu5-`Eb8Fw0O}OGBJW}dI4BZo!B7j6>qLsO#pdKek(RaufwuYxRx#KLFI|O9L|t4GmuBpSu+?AZXJmh^n^~ z1__h#tUc>04sql|Ws61?X|*cOPf#|-B`na;cew@F4Ygg>Opj)!{FjdQJhcc#-GEOi zhTM>0DYzo^)f9aF7W0uu8uW`hWWQHOKaF-)6kf?wX^c>U z<&8GEtsJEBv&-?mW?+__`Ro9yVMwW+IG~!<$r@ucn^`VZC_~Y>=|yp7XMu8<;-%2i z|BAl7Ph6%=tWU}1h(l8~Gl{B z@MqlcK5m3mHn!MTV~w>VMJ5+)wPmj*w)!*IExvVgUfEs~)LzpRxgb>ZhACj&63_0j zUtW)E@qdA)KJflgmoFM)KC)-FGXBtpdjE!=IO0<9o``P|*2#xhrc|vy?S~62DzuGL z3dR;h8$;~or8v3ShKI8Y1?MSDyiTfrqXAWGEP5G>6+9Oy=heoXaAXPJ#QxMdXBldI z?2ndmZhI|rr7}=1*l)-g!;?t`GXLc?S<+d$e##+@oG9*3L)g8UF^6W$Zn(-j#=1FI zt1VTvAPRvkir7|=4v+&UJiooQo75?ByhTPF{ zFQ>I9hauRnrC(&6zMmxWpRi3pi#;@-f@b@TsS6Gau%Ot8`psEJ8OlKacP1G|4)a`n z*F_jPY{uu9=O$VXEg3d0$<5CU#z)&zPnp@NM^_!tIcp$88xR)D(xJOb?lqB876^r+ z;4dQ0g$3`|vtp53o$lotj<=@IIL;H}=i02H$e` zv#xhb>hfJGct-|?Y_-OC0rkvlk)*eEm03-S*KQcHAu&hIQY3@~WQxA^RyxZxr6RRH zT&1t;Mp>(K*0o?#&S8rQnA>jjWZ3Ys=g4Kht~vJ>@sz{L`iKmT%GwtNjAgc#ydzXf zar{li*TLjX7&uhuUd0mxsOp?92(>4oDXEO526$H*&!JHVmZ@ZC(Mk+gykLoa*;(pM zB7qaBMQf-~ZBx`ed=Tuov-^cXXAnokWZ8GfPmDDcpC zo0^JPvG#2os`|=k=G`xAnpLa8JqZG7cuT?@D9mQ z>m%`2P*d|{St)#Da<zj z=-#DrDOLKyd^p|j*Hym3=uLi~oasBqdrHHgR;~EsR|N9p$v<>!T=#~Rit2oPU-c7z zyzcF(DgCOG=YnTN@2Z{HX$H+&=9UN|uE3u;z0?;M!AmjIQjFm}WOS1wqYgx@Q5sEk z`3bN_G2q^~e1=AcILID+-e(0TZ8)_bX!Q|R3=hws_PGa&R<;+;KpQj+Zi6e|HnV7A zM~CIZd-g6m>EPthI{7Kp`q}DteZ4@V?RkevEw9{|%Xw)liT(v^?+8iII*FR;ao+Qc zG}W0XrJ%=%H-XuAuI3@O$Cf{Q%*x9nMt9*eJ5DsSeSEs6{0dSF_$|1)-~OnLLOWo9 z`XcMu6>GC_-3xgt>G3tUFW}8Zcl7atpe*UUMn^HT^$>f=QP7#@%C9rB-`2QLjhK!{ zt*0c?n8x)o!59Au)7ZD)Q4KPiTJZ8ct^`_aZ&Mi~oiouggt}y%r$#t^OZ{v7_+Q$) zA-9lPuCMGgIE0TfX&}3!ltbtY*(~{6XY|Pod8cK~IjvW_4ZnkqJazcxjK~&Kv7}{I zDzQYp{Hi7&gjXu${>}WI@DpLu+K!IYGk18WC6P)fNeaF{%W4aOLL@e}Hn}5WDL8Q;B39oR2x@jymz&>m4Kd z;ZLe&e=ANE-)oeIPLvdUUis{mSn^ggrGId(zo4?LlHRh?rIY7A=6P$~V!iw8te%!t zvsS0fGW>e?Imrzvt4ZHu8u_=l1P` zKRI-ndw-p&14uOVRMYhpr>DknXOHm3`6=iJr+4XH1X_xm;G(?^S2)dQ*Ij`&45|2` zKSEny2IMj9SFrIAmr8eNw~QwAEV-!7Gm*3lH0dv|RYKvGy#N*Vm@IcVXCd`hmL*8v z_M(Pr&pd`T{1yL#1hq5M(p@HmdG%1sS!~7Sd&l@IJ>@Y}-z@(vR-R(ldqqHtBObc@ z!d+L#+^DTE8qwC8EhL8;j{G41pACka z6aKcSZ!c9^2VZ$1k{|T(_;&DR%7)w#%>(%++t)d0{WQw&ow7a#bU)p@v^17sR<}3p zD3{uO)iIeZXgGK*n&fnDGAjgL6tE-Y6<{a}tPPJ{bT}hAILwyhT}JQU)rqO;Vpn#Z z(-CIG*L1>qCbQY$E&+0URPl+GsyR8S7n2~+9MLI+Htmp<{c=r^)2{uJq*B4o#2U8o z=Q~pMuBeq8%q;ododjIE&CfVuEVq$b-IeBO@D5d%Mz*&M7CYUt5NXyXu~YPws!m3l zWuKr#{P}&J;We4f`ghWW?1PT-xjw-?7C9ii3!;Z<_y#X+yW8}N=qgT^N7I}?eF?K@ z@Nkk8nZ8Z2{yXgpf|H$J0i$Dy;K(brQlt!Sav*IZH;JR)RFnf zc9<$}PStMRb(bt!xA1IAYk1+>$}|Txi*!Mg$pdwcDfB$aTH1)){%w)w;&qG-pY6ovmz*SB4-6mtdGzw>IzM#C5-v!Gv z>{ch%qFlI3T>pwtiK~{Te*H5{!=z>eGIRFt<&1B&CC)iEo!vvZv}5nKR%$TWsTAF8)X67QDJk$C)CV3v`Y^D~s))ngvM>fH5I9Gj)q3>H_+V==!aKy%bA zf0!PycQsw=?#XtUJ-CJq_BvYYKQI)-^JYC&Hq1$fq#ld=EokcIPh%~CwL+e5<;IbN z7hD2SKBon?pI`?|V&stp%%Cpxss5EPT21w%G5yUBB^woVhO7^xo}QJ9$Fz%@XXhu@ zs=tV-XKXvRU+0O-GXbz8C1Sx+xhByCZW!6}maWxZS_~29=ejUa@3*Xxv3U)s969qo zo|0gaN=!PZ$i8t@l$5T2xXUYcITBfeaXv~WJL0to@6ei={Ql~T&aLC-Ne0W6+UVL? zlc$t=hKc4L-;CJQ9UFoluu-0yt6TVOxSkuKi7Q@XDX;dPEQMz_C2?v)=m$r^R&QeT zSFfRemczqqUD$unZDks>4?Y0Yzkt|CgPo2}vw^dsxA=@%z=Js3?|S&SiRbt&@<&ex z$PjSaTfQx@xWH?$O0=ANP7SUZn-;CC%*Su*^O2O_|I2e3qlk(K*q}A6{G8i5ZTS_Nwh1YXtYOk>NAJ2OU~=DnIyWrK<+`D3)#!$c7 zfuOr15MPER$*^#_MhU@xaebQMxQVXn{W09;ze75Qzse(>HICY>zD6)wsYtOn1r6w9 zS}Li5^bmcNe1@T^{zWh7i)ZZQzE7)|DUr`Mu!RCaSblX=rio}|2_r*IC}!kuw_DJu z%ydOff@XK`bkCl3WQ@9NlcJiKrf(hAla!`kI(y1k+fSxysC}?)L($0xwLv~@W)LtT zp)bhixXnMmK=;Pueb|1~yTHM4H0t%!rOT@ou^(^uRgVGt2M`xN>F^8vBY;#E1M6nh zUJt2)X7=sfID^(XXE!4_QkJyf;ANEOWqo{*bXHuJ-}x{;0M_GoE=8dD7v8*)az|V} z2Jf>rHMzFa4a7caN>QX4t0J~1)D^Gx-^Wo5l<*`uYlwk(ZJhZe=&BAewio+HpcLBB zNQOT%2t(l04Jn3Yn(&qd`{?GlCYN>JXHmUpuXUmV?t^rqcQ%`g1Zd`=N=f2GiM*26_PKvlyt6*)Oa#cXjdPjxQv z1A+vUN3>9!pxv5c)W{lAjLcPx(X-k8c(PW5O8*!Y+C;iWpjC`)Wy(flDr6wq=O_Xc zf&Oev;%dPDTg_KKBE2&EOf0KKd4yQygIGcCq{GOW^0@mL)pi{X#6ldHFMmUsy^p!A zZJZun-aesyS8h=$X<$%38LNKxhNkUm@6BeHVL z+1%C$a6=#3+E4*63nN7Fw9x|`t+s&$*S-WSHOT!x4W0QzQh67~@8!x8vRRr}1%Yfa zHN_owkTkUo&D6>YLd`MX!X?eg3`D^c6T%DbaByHg-eG;zyfpMyPACC5ne6cddIg#8s zpCsJzjxH;NOoLpVxSGUc;wrz?ZE7@wb#?n;A(>HUslXhW-X8B*uCYZkmn+??TmfU4 z>;*R=P8WqH?rFspw>2(4v7iaux=M*kHT|1+I^#B4ySuOMy|5vU_x|NWt9+rYY4b1N--z-kPp%)2Sv)B=%|_jL?6g^vNx7xQzgQowrAqk9M6

D{m8vwDZ;}%>Yzr{o3w5@5 zg&hk$tPe(SH_#}VpQ5X)42cn3?yR$gqJyG#ubfP2I(Yw1U4?*5a4-^JA_@K>^9`q! zo5}DbT9bYPK}dDCGTgKwMN*HkFp1{g&q%rUG55QTUI$SRGSyWD|Y)r@S{0X zw{m2^6UFbFr|i3}jW4?W->Q}u7_b0WY2B_F96&mXz5s|Lh0d-OhvK-afac~Lg6Wfv z--IoSw;p6KaYubUDRQIVA#Lb|F}2865s@sgm6I6mm}Y?&!H@wHf1Nl zQe%QJ&*N%8rHJ=qb%X`$YyPJJ`qS@x=|pYZ`R@_d_JJYSM>FRe5PeY{&7-az`)rK4 z*``7VGzE2A9XRLocc{Xc9#p+DMKVenen}Yh^$e!cUJzW$J$7d-f+tL*8$J_2bF~A3 z=$W_o*)_0e$4t1!A9_vgg=dAcQn|^v=5{g>gfYAFf3xMtLw8;cbAGEbXe*Y6o>6cip*Xe}HjzJ0mgWZK}5hnmcK=72%uL8?!&r zTG&3bbG~&|D2q-eyGk~j;45;(ZF@z=mbQZgMtXnix+78xQ zCH(2tw8PHB?#%U5JZw(MDGI`6{*!0d{ya$jcD*$GC6Kc7orO&_0%-1ko4i>07*GAS zjv3aUND6e~rP&dM`IP3bU*zz=pBUI-MDFW z_E!+i%CIR8{V17WEv~@{$h~AT93n4|A5LH%asijKAC^$-wFuekwE6od3zTCE4_;lp zeg3~aDvzT4IX4*#{f4J%bW>L*O4M|?f_3>V zM}!@$V&iqW4ycg8xePUfv=)+xV-Hj*-?W!3ueO_f!+$4yzr4o3yEp7sLPj-~{C5$| z6xtm%U8sq{e&GHDY8Mh;ZSKU=2v&GuFUKN9=~(GSnX}=O5o60gRzi7tT=M z2-Tiw%&Kcf!MA08P)Pf=8Q)(w|M(lYP9yuRS;N|^wQQ!~jbW?v%r^%9w5P!6D<$(7 z&Y!#10JF7Qy=>;x+d4MP*t+h>+_ItlruB(7P$8*#!HC`NmhQ!Kq5i9{uD3JXpHdsx zc71REb7|EEl$;yl-CjkE*QHY%s(uTX?|}*)tBWgE?veCix05YIaxN zo|}r)doFgb{s_ErtX+7#^hM?_gdgo{#R483J-z#g3Bb7PHUPOs-wf_ocu@P^t8+f* zz@5b(CW`=3`m3F>a*|8H1LQIM;m4zUg98D*OahqhJzj)p&Y&ZZ~Z1u z8*i-A^ef4;gN3nOgNDTSSM6-y0_m@-vpPhV0`fI}1dbxYhtOZ(9GPg54a$;5-kZ?s zR-e5??q@Q+W4dCMhK*l1qV||}k{XmFx>>%tY}&j6u;`)`(0xwZ#;qWXCFWmb|7|H3J>ZjF;)V1TC{ z@vFN_);3MZ?Bl)J^L7Fgwx-8MaD>O)>cEc3}IPIpcoP<(3Q24HCa>K{*<9vt$#;J!pTxCO)@T0z1zjn{Z2h zxoXfL|9w+P7p+R8;dyZ^;NFsJKSVRUzR_j^Takapt?PIXX{KVetnuxUYUaYISt@_p zXe64(wY5)i^~5QAT;e}A{SIKC9x%=8Zuvc=vueT2l7r=8E?c}Ln3LgF3m;`J4hE!q zliJVU+H;F*8XFS4Xhkt2Ru7T0lDn>b2SmHK&0mAFmQ7~ABaOtNzdG`3%H1p#da0+A1iF_WcuM0I>M}QQ#V^h-gxwYH`=K!KfORJ z(=#)gHD7Jca|c?XizAn^T8Ou}F_qMLv$^<$5f2(G>!}Dn`o)1?bNz&0V+(NJYZtXZ zAAVVHRDdb`A&0YvKY*kkyGjA?aY0+=6=P>G=;M2^>yK}!Fid$!|239LI_u#=ywmkK z%pu`{nc8hj#EJV|S(^iv7ul!+dPv)%@@+#TmKhR~`P}%cR=NI7brcwun{{Oja9aI{ zQdwEhyy}|B@%b;i%J_*E3FQ^IMYG`u5OVIoziWf%+OJ|8Q`7I%2YJ%iWhgs;Hc{cd zV&=cemb9N`TbzAnD7l55r5Jwu{0iSv=N{f(xOnIpf~8TfuRI zd8m;n>cuslN(0ri?f+g!!!ce`Bt+%QY##|ajRB9k{kHv~C}z=Ow2ZXfu{zc)T$7N5 z!1$_k7x1`1>z28%`Q zQ^8_jSNC_wfSSyH-h>xE32^m}^^?@;GWt9zU*)&IHw6Jm@i|jsOdmccD@8zn1vqA# zR$2s4Fm!|$xjyDG^3h;4!f}Xr&Y>}yEte72n;Z15f$`;Oe_D$Pv*RXXpIk6lQ_N)7 zUCpOm!(Dm=9Mz*GcFS2;t92Zk)2a3Ua(RAC;db2YMbmLTRU)79gw+u|pM^U)cSZKJ z3~qXVXJ#*~qZ*I944}bU`V@LGmDEI0TcX+?5cIrlqiYCyek~@sy(&^=^w41fLrx(@ zP5&8_S7`*#a|Simb^DJ{Fb|P<`Bt27kwu2V_v@-ZgrnEGesP`sE_y+{UBk)aS=4@w zM!2c1)NJl`)@lj}BSC1_F?hm3*|(s_WO~1kt8$*3mKJfb7$vdiHaQQYLjvc`zjM?6 zWT0oujpa(^+K3m%7#a@3E_I9%>?15|ZrRi^dCdP!^?2?bj!GjGL~ZF}dDRgt7h>>M zek6OjOk1~suySk_c?HgUJoYHn=$K^Q_8{|q6_g}m9d6kfxongLU zOIRvyLpVJG#w#w*tRv%QeakLqKOD24WD|Xdi=DGFjicPYza&G2PPWf2c><45dydWO zol)3n9sx_xg**Vsa=eZFhWrWKW396-sWkyXycFoo0!wxt_c#kV0geTBNF+Xc&Ap7? zbc-8AzZ)E2!aEfAn!I(ASCA?nlF+5nri}Re;{pPps8*KB?h z$PEo5Gn#&%%`Gmg*_vu2LqmOZb@==ah+cxO>6^tb zXA#hu#mlU%bZW#m$n;MaYePwYV*jz4bUu6cZO1VmxWgO5X$6BS#R$0)UZo&!AFg-;MFzbv?D z=v%*X9oEmIp}ad67g_52G#7Q6pB5wxcNyBXm184sPoAX z<3F#_+yKwtxfg61l6{`taF4i?8L6OO${H%g?wx1?6Zo+MzMSE_{WXCejU0x%4q6yr z<`Vs~lQ3%vo(x3dT&-;vL9*NFT8l8G)oW$QdwOpfBq8&*2HNg5=mQDS)OZ<0(3bsT zPR-k7{2`I}i;@7nZze?&lA)<%?wFopu2}ddkm!eK!pE3J>10GI%1kK)E*OI%)Y1IBCqi@)tTOYn? zCdwJI+1OfB36|xaK3TcH#Toa71o!o{&hs#oh-8XySe!t*L4`c30enzX8vF}e1_J??rL4A4|_&f66_Yl=tHCDd&_sKv;&q# z_kJg|p<)U~`O|i1m^XyTHiMtYw(tP%zxh3`{k4oo@v@v~XSHUBX|^$g)3DRNxBAg?$T4)bJ;N`Xt2qy9M;SwtJ%m4g%nmf=ywU@>eA1Zx)dYCM zlut0Opz0^Em{TV&CZEMwW(j5qI*n|p-U2dS=*m>;>ZD|!U+CHxeU^^WFwKdAsJALi zAM8!d16h7N`aYjP{6-UR%=Osg0bCCu_23Yc1227uQ>uLFx-;b8NMrJdarEc;-Zuvj zO{4t^P=?J=74~d7Tw7{>Sz+1z3_%C9LoD7jIL(oO!VP5THt`KkCobQf_Qhpj_72EJ zC@HloML$?Ampx2L9+~Oh$5&duFX+}D9fmFSClH({1M3Y1dz*L2)l`ysR9XPPYCn@r z{N;8w^NCl>ti6pYA;-teTZm%vymjQ?Z(BpC5EbJR*>j(o>)n95)6X3vhZ;u6p$K!A zj?+!hnZ9Ac3c)qo5T#^>ykUI|$NAf*A+cTEul3RAA{K`d61awX-VF~SxLrSjQFLXomXv@os8!80F-~eAE z%;;*i@9g^;%{X|QN9(*C7T=JEVurHij4?9&l{Op*NruR^K8%|ipqr7T3@XDtU{~uLhk3*8@>H{MuynVaD(swyb+9h9q#JOk&yc@?OqD>(P#AwT;@-(dJ~w z0Y~6g`ah1q7|fk0b7_2`STEYrunFtB$CXre>Cp+qLB=Fx?UL6PUr&1>_$FzaqWPse zSi#XR*U$wP--!+(`|ClO{hv_{XG=TIiHJRuKB8@(FBs0x#JmG*|pQLvGKVp@v# z10R1>G`nXA&AWfU+xq~isEG)!63RYh9m68T4QU}wryP7P-8TUCyZvWj1V4eONpl^ zY8p2dFBhUUf%5>r06PW$U)A@Yz@as>dDI8mPtFEh)8~)Eu3$+3H~bL$qG%FZ)`_jW z%kNdgV@)2RGG0DfsDnnRKDoJD$?$-^5$2`ImAjUpA7+v)|Em-^m z(NVoh^vEeE1_(W8#mu#=XyOlr4^E?Ef4tVFcz*s6qS4i((ly?<4pr>t zIK*l@c{&h>@(EnPR}FpuUbRiP1E^GWNf(YBzUtw-nbuW;&<{br94bjLo4N`Ehf2vf zyJH%mY51juH*`qCX?8~`Bg&}vCv0sgU4(TtEqq8@CnOId)uai{- zglTpQMCR3F%y7jNVGMj`O$LAMpH1H-M0Nv_JHD*npoF>qTb>ClDzq>Try9Bl(^k7g zT{DXT@IyGN;e$h`%US7HOyWRDBdA%jaxoI+tvBC^?T;^R$PJ3Db~+fY$dG|1#{B`9 zZc79+!n5LtQ8#Nye&$?iO%F ztBo0b{DKBe@3i{>WRa3lt`E&)aJ-$_=M#mS9h4=JjD)Hk$z(ZL(+Zm4j8%1t`-(|X z+9b|9=0I@<3l;vZl$G?f`G$(+f2QSb!Xs*9w_2MMzU!D6xfyI3Xb{{0vAM8LF(v0~ zH_Zk=!y!`Wo3mEj64G#M0fI8`BK;FpV^aGQ+T`I{1NUYIs=T}Hsc1BOSvk6UXVh1D zI8?2Bj)w$;tq|il4(1QMa3KH}DvJ$nv%~MQ*VHTP*A4o0zSNu~tpD>9BK_~J)MoF8 z)frK3&3VaVG@Ntf^B)7c(?-T-AG@;6(0OL4(jQ-CaKBnZhI_p&Rfm6bt!2HKon0~< zL>r73pXLt;i-A828MgcBe<#ba{o8F1()UwsK6?;Wq)fW;lFK4wK$qBHPDfDzvMg;6 zw1F1ZKUFe@QJuZd@!ZJ^FdQP%ImIyL;V6(WdwR`wih}lY);s zn&dBhkLt0GoceU9zh(u&fz)Rzh$d0Qn!HXL;ob0nzqN=gB~ZG;C87%fnVb9^E=Egl zdu&Z!v93IJy(dA_hzGzKu+R}+MSgSl{AUP_Tk!m=?myT3=FOXTnhZRy3@2d9ZgPo?g(nwwpkyDJlIc?lYq&hwd|8ESCGiBWH?q8DoF;Qg zmvFOUvbkw8yzb18{7Y_89b2jkmHg!j$A9$g>Cmgd2_9UOLYAOb+>XmAlyw8IHm;@1EX&F+`qeQ6}*YHDEbO>1Oc!{Zen!Ys;A*QFpwiL#JLsZ|OtLh4;Ku|{VW z>TW=myQq3K%6xy6YaHnin2I6vCAojMr6(YY1d7qIh&H+0=4Bqef4dh>lj(i7p4g&8 z`A9~XL^FoW#kBOjfzM0*S!+9O=d{e+s4>#MSv6PMsH z=avx3{o!?-gBgfP17PFU`x^FXkp(2$!e@HAOxLJtGs~Ypz8^Y zQJONPDZCd_(Fz^&NEQ;_9vDWqvRuiMG2H2R&G3|V8h6uCscQ+qW$(_A9kIYq$$5av zY5T@l_uIa4fj1PyQyKX)LHs&%)IS5jx?===E4X%~^wgcwml!x<;`k%9ab*U0K?*!c zx@>HW0&ZUT8!%u8Yzsuf^Le{@R+O=q;MTdz$3y{FE+$=oUbXF}CECymSP56w9#vt> z2CW}KqCUvFi8`=BYle$C8pb5}iDDRL%aolnORS|cRg!&R_1Z*tZB;!qV?Gt|a7PM< z!+F%YTXgxRK1Aaert`S5HPVnOpWBLSPMp6{1oAdH>GMDQ{u9Nv zw7LY54i?xYIGpY34ze&H9bk)lWknE@{$*q7j8~$5^e5keXf(j+!yiT)a>^J1bRi*i zG&Py7gYTdAOMqd9j=o9P72lAx0+!@*iM&cQssflza#WK61Zt1O$9N z$0wYg#?U`HiIngYJ(ERdEoyL0cS;=xcJ2~o^oF5As!^03qL2pFIrF(r8WphXjoQjM zIFm6N;J--t2zB-CiD_C^pyp#*s1!EuP9Lc+;?y-dfXXx<6Va7RUq<&+h!Y!SyErYSxs!H?L zT^-(6J=vWLG={FI;S87|e4M3rgeBbsixrMB60nOt>iHW8vJWxRH3JT9{V9Gy9R4fG zcsQ_?C?E_GtjU&BDvZ|<{pi7PmJeJs43#ztTg}e&{Ra$ouGIa9yh43)&h&PGUvnKP z4QFikk(W8hAhkeM5OasfHt|B&FLEp;j9ErK2c%W9mGcgG4N)B1i71tPfeY|cv;q{R zyx%fxqi`kN;q!LN&rwQcb6|&U6>A;XO5IiLZ=1|l!PqiSPQorh?&H=4$D$;{VVzjc zEi${HrQy=m;3^J9XHVbcBVvEHlSD6hU@aOk^zx7FBl@@vdlQnUP77zTtvpe(2tlg-iMd4tu00Sxe2`=gMKPHS ziiDzYI9ryTYZ?wgg($!Hvt>YyS>-&0%3{Aq=Z8ZsfH^mju6S8f5tI+dYUNBSe_vg* zU%EXy%_8UEV0I3n>w~wNmK5_1X2(k~uC^+(%R5I>2G=8Ifax^E*lSPXF%0HHRH&9Q z3g^yejCl`a=nOvr?&GY(JiRkkbGx!2U}q)1x-@Wpgm~)o^=J#sPkPpicl4=-ty_kR z(fwC(A(pYt@_-bfQ4kvXI0D#Pc6Yf`w{WFHmb?P1>@9OlSs2beyU1l~M>cPEblFza zSIhuCC38T^&>b0Zp;zKU3mhK6(NLXc0^3i`iH91r2-{VS?e=AihW)<*7PSq^mFW?^ zz5bdp3I~qZ;e(@J0X%Ra^bNFcR$Sb+3cq0V0O+CbA`D(kNWro~lsvA+!r8fG)a<=< zz9kjqInTf+MWY5F1nZjR|lrN$o`8xx2w0tTm#@+lM(4!p6fl6~PL2 zZRq;wNm;%_3jcFmH7-QTFjLFOIJW`{$uP|T9M-W{%hF^b?hf_|O(LySouHg3vnk6` zL=LSHb#kkP2`>k{mF5wgg$p)8+49|qe3+X+7Na!~+Efg%V>Ia%X}oOxIXj)2r54B9 zjcpNk3eM8X*B-0Z<`Thiz8n=mWou=$9Y`%Y=K(4eNClr+rR+BNh#f3{hkXaMqDy2M zj@}?w+maofX``bBJQq-M5^r&F2U$M|w^@K~4*#~5S$ zhgD93HAwzqbo4VT{>*E!=`gFeVLVLYYJRU-V`~?-0$5mlY(wdxCluEyG2v#?Zt_#BU|^HX))M zA3f}aN*LrD1DLSsjj&fU8YXomGYI1MR$3UsU|*G=#n|B$+-y4uX1A-R82`K_W4P-P z_GlGZ!!hZGiINe<_jg=5=glkn>#NVx<;h^&=0#~=4k?pwso)K*-0bR36tO+fIdm}r` zmnPvzg^oTU_6mM^o*Nw63$=kF3C3PQbH!TH#?+MCKK7_`44ds#{eDNa&Z6bu*^ddc zR3#`ESNB@w1dPUfL@rZ<>p`IFz7$2Glrx=Bj;cYqBc(Tiw3Yk^hhJnS&KazM6jjh@ z$lP8JwmdEE5a%|fyduOcdn{C`tzDpIXo@T<;C_i}y_jp4#eg;Lm514}GZHnc;Y;pF zkkATnE?(MUOEM|%4{Er@qVRUWtRUAgiAA^4d3nL%EEyKIebj^KvXdZ^l%i^dD2)_qW{>JgY)(Zv-80V<^685MW5|7)J z3(bsC>Hxh$lVbg}JQGIFr48Vaj$Y2=$%GhFslvq9JhI6pXg`}_i_>XRWCa(pB%)OO zCNPt2az}j4Yq4PbN%pLTZsODmKck6)M6|X+7s|oo=L_Of_zLr=KH*T9>~G$KlixK; zBw!bzK~L)8jhUj8Wxs`HBj2xFnSTT^nNat=bw5~F)<#j0-00^LxM7Wr=i2suw03$Q z!fIBXM8@qWF-HHunw`Y88r^kQs=~yZlJ=MwVMy*SaHqgwJ)#i;D~V-b`C7D)&nw~X zirOndW^_BR09+68)Yy()GErx@>kJKdYoL{@ZIuCIDX4hHgIVfaX9&jarfW1`tOy4j zOw^oG^d!m2-!|Tu8GO7Ys-&uHNUokRvZTm@ad_kTPaDS3la*{E38-*OAv{IHb>d1) zmc+Ade7j?cqhWR-ZszvL4aUVCNKf;7E`Z(*G7%b`rY(&E!y=y5qRYE9*r#1$;9g5a zlsUQez@TJGvoX&Xc}Diexn?v0wZ7`bIB)zOKz*@CaRi-kKE5K5SlKkIeMPI8_&uf} zIsIBB@o?etG9WUmOoOex#c@rMgON|lG$Uy#!3Od^rMLW-Tx$(Dkgu)(<1znY#V|%l zKfGOmbsM&Fk1hYDiU%Z7w+LetrMQ(d0fbz-b#z6pg zI@4vkQHF83l_Zt0C z){>2X0W?X{2qHMhwz9!PY4t*!>iZ7E`@O|9@|@QWhO?5DSb*^F5=GwgrfJ-7KOg@e DASC13 From f08302e2d9736e02bafae73c18f977c1fdba90af Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 13 Mar 2024 16:04:47 +0100 Subject: [PATCH 25/43] fix: compilation after merge --- .../{V2/Lang => Composition}/FeatureAlgebra.agda | 5 ++++- src/{Framework/V2 => }/Lang/FST.agda | 14 +++++++------- src/Main.agda | 2 +- .../Experiments}/FST-Experiments.agda | 7 ++++--- 4 files changed, 16 insertions(+), 12 deletions(-) rename src/Framework/{V2/Lang => Composition}/FeatureAlgebra.agda (97%) rename src/{Framework/V2 => }/Lang/FST.agda (97%) rename src/{Framework/V2/Experiment => Test/Experiments}/FST-Experiments.agda (95%) diff --git a/src/Framework/V2/Lang/FeatureAlgebra.agda b/src/Framework/Composition/FeatureAlgebra.agda similarity index 97% rename from src/Framework/V2/Lang/FeatureAlgebra.agda rename to src/Framework/Composition/FeatureAlgebra.agda index 12911b05..7fae83c6 100644 --- a/src/Framework/V2/Lang/FeatureAlgebra.agda +++ b/src/Framework/Composition/FeatureAlgebra.agda @@ -1,4 +1,7 @@ -module Framework.V2.Lang.FeatureAlgebra where +{- +This module implements the feature algebra by Apel et al. +-} +module Framework.Composition.FeatureAlgebra where open import Data.Product using (proj₁; proj₂; _×_; _,_) open import Algebra.Structures using (IsMonoid) diff --git a/src/Framework/V2/Lang/FST.agda b/src/Lang/FST.agda similarity index 97% rename from src/Framework/V2/Lang/FST.agda rename to src/Lang/FST.agda index 29483430..e6d8729e 100644 --- a/src/Framework/V2/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -1,6 +1,7 @@ {-# OPTIONS --allow-unsolved-metas #-} +{-# OPTIONS --sized-types #-} -module Framework.V2.Lang.FST where +module Lang.FST where open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat; reverse) @@ -17,14 +18,13 @@ open import Relation.Nullary.Decidable using (yes; no; _because_; False) open import Relation.Binary using (DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) -open import Framework.V2.Definitions -open import Framework.V2.Variants using (artifact) -open import Framework.V2.Annotation.Name using (Name) -open import Framework.V2.Constructs.Artifact -open import Framework.V2.Lang.FeatureAlgebra +open import Framework.Definitions +open import Framework.Annotation.Name using (Name) +open import Framework.Composition.FeatureAlgebra +open import Construct.Artifact Conf : (N : 𝔽) → Set -Conf N = Config N Bool +Conf N = N → Bool module TODO-MOVE-TO-AUX-OR-USE-STL where ≠-sym : ∀ {ℓ} {A : Set ℓ} (a b : A) diff --git a/src/Main.agda b/src/Main.agda index f952140e..5ea4b6fb 100644 --- a/src/Main.agda +++ b/src/Main.agda @@ -25,7 +25,7 @@ open import Test.Experiments.CCC-to-2CC open import Test.Experiments.OC-to-2CC open import Translation.Experiments.Choice-to-2Choice-Experiment using (exp; all-ex) -import Framework.V2.Experiment.FST-Experiments as FSTs +import Test.Experiments.FST-Experiments as FSTs open FSTs.Java.Calculator using (toy-calculator-experiment; ex-all) {-| diff --git a/src/Framework/V2/Experiment/FST-Experiments.agda b/src/Test/Experiments/FST-Experiments.agda similarity index 95% rename from src/Framework/V2/Experiment/FST-Experiments.agda rename to src/Test/Experiments/FST-Experiments.agda index bc57c2ab..38beb57f 100644 --- a/src/Framework/V2/Experiment/FST-Experiments.agda +++ b/src/Test/Experiments/FST-Experiments.agda @@ -1,7 +1,8 @@ {-# OPTIONS --allow-unsolved-metas #-} +{-# OPTIONS --sized-types #-} -open import Framework.V2.Definitions -module Framework.V2.Experiment.FST-Experiments where +open import Framework.Definitions +module Test.Experiments.FST-Experiments where open import Data.Bool using (true; false) open import Data.List using (List; _∷_; []; map; [_]) @@ -20,7 +21,7 @@ open import Show.Lines open import Util.ShowHelpers open import Data.String using (String; _<+>_; _++_) renaming (_≟_ to _≟ˢ_) -open import Framework.V2.Lang.FST as FSTModule using (Conf) +open import Lang.FST as FSTModule using (Conf) module _ {A : 𝔸} (_≟_ : DecidableEquality A) where open FSTModule.Defs {A} From 9a8d26a040b74d8595087a29ecfff2f784022950 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Thu, 14 Mar 2024 19:23:58 +0100 Subject: [PATCH 26/43] many refactorings and refinements in FST file --- src/Lang/FST.agda | 568 ++++++++++++++++++++++++---------------------- 1 file changed, 297 insertions(+), 271 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index e6d8729e..2c70e4d9 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -1,7 +1,8 @@ {-# OPTIONS --allow-unsolved-metas #-} {-# OPTIONS --sized-types #-} -module Lang.FST where +open import Framework.Definitions +module Lang.FST (F : 𝔽) where open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat; reverse) @@ -12,19 +13,20 @@ open import Data.Empty using (⊥; ⊥-elim) open import Data.Unit using (tt) open import Function using (_∘_) open import Level using (0ℓ) +open import Size using (∞) open import Relation.Nullary.Negation using (¬_) open import Relation.Nullary.Decidable using (yes; no; _because_; False) open import Relation.Binary using (DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) -open import Framework.Definitions open import Framework.Annotation.Name using (Name) +open import Framework.Variants using (Rose; rose) open import Framework.Composition.FeatureAlgebra open import Construct.Artifact -Conf : (N : 𝔽) → Set -Conf N = N → Bool +Conf : Set +Conf = F → Bool module TODO-MOVE-TO-AUX-OR-USE-STL where ≠-sym : ∀ {ℓ} {A : Set ℓ} (a b : A) @@ -48,283 +50,307 @@ module TODO-MOVE-TO-AUX-OR-USE-STL where ... | no ¬p = ≠→False _≟_ (≠-sym a b ¬p) open TODO-MOVE-TO-AUX-OR-USE-STL -module Defs {A : 𝔸} where - data PlainFST : Set where - pnode : A → List PlainFST → PlainFST - - PlainFST-induction : {B : Set} → (A → List B → B) → PlainFST → B - PlainFST-induction {B} f n = go n [] where - go : PlainFST → List B → B - go (pnode a []) bs = f a (reverse bs) - go (pnode a (c ∷ cs)) bs = go (pnode a cs) (go c [] ∷ bs) - - -- the syntax used in the paper for paths - infixr 5 _._ - _._ : A → (cs : List PlainFST) → List PlainFST - a . cs = pnode a cs ∷ [] - - -- helper function when branching in paths - branches : List (List PlainFST) → List PlainFST - branches = concat +data FST : 𝔼 where + pnode : ∀ {A : 𝔸} → A → List (FST A) → FST A + +induction : ∀ {A : 𝔸} {B : Set} → (A → List B → B) → FST A → B +induction {A} {B} f n = go n [] where + go : FST A → List B → B + go (pnode a []) bs = f a (reverse bs) + go (pnode a (c ∷ cs)) bs = go (pnode a cs) (go c [] ∷ bs) + +toVariant : ∀ {A} → FST A → Rose ∞ A +toVariant {A} = induction step + module toVariant-implementation where + step : A → List (Rose ∞ A) → Rose ∞ A + step a cs = rose (a -< cs >-) + +-- the syntax used in the paper for paths +infixr 5 _._ +_._ : ∀ {A} → A → (cs : List (FST A)) → List (FST A) +a . cs = pnode a cs ∷ [] + +-- helper function when branching in paths +branches : ∀ {A} → List (List (FST A)) → List (FST A) +branches = concat + +mutual + infix 4 _+_⟶_ + data _+_⟶_ : ∀ {A} → FST A → List (FST A) → List (FST A) → Set where + base : ∀ {A} {l : FST A} + --------------- + → l + [] ⟶ l ∷ [] + + merge : ∀ {A} {a : A} {as bs rs cs : List (FST A)} + → as + bs ↝ cs + ---------------------------------------------- + → pnode a as + pnode a bs ∷ rs ⟶ pnode a cs ∷ rs + + skip : ∀ {A} {a b : A} {as bs rs cs : List (FST A)} + → ¬ (a ≡ b) + → pnode a as + rs ⟶ cs + ---------------------------------------------- + → pnode a as + pnode b bs ∷ rs ⟶ pnode b bs ∷ cs + + -- This is basically just a fold on lists. Maybe we can simplify it accordingly. + infix 4 _+_↝_ + data _+_↝_ : ∀ {A} → List (FST A) → List (FST A) → List (FST A) → Set where + impose-nothing : ∀ {A} {rs : List (FST A)} + → [] + rs ↝ rs + + impose-step : ∀ {A} {l : FST A} {ls rs e e' : List (FST A)} + → l + rs ⟶ e' + → ls + e' ↝ e + ---------------- + → l ∷ ls + rs ↝ e + +mutual + ↝-deterministic : ∀ {A} {fs rs e e' : List (FST A)} + → fs + rs ↝ e + → fs + rs ↝ e' + → e ≡ e' + ↝-deterministic impose-nothing impose-nothing = refl + ↝-deterministic (impose-step ⟶x ↝x) (impose-step ⟶y ↝y) rewrite ⟶-deterministic ⟶x ⟶y | ↝-deterministic ↝x ↝y = refl + + ⟶-deterministic : ∀ {A} {f : FST A} {rs e e' : List (FST A)} + → f + rs ⟶ e + → f + rs ⟶ e' + → e ≡ e' + ⟶-deterministic base base = refl + ⟶-deterministic (merge x) (merge y) rewrite ↝-deterministic x y = refl + ⟶-deterministic (merge x) (skip a≠a y) = ⊥-elim (a≠a refl) + ⟶-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) + ⟶-deterministic (skip neq x) (skip neq' y) rewrite ⟶-deterministic x y = refl + +↝-return : ∀ {A} {e ls rs : List (FST A)} + → ls + rs ↝ e + → ∃[ e ] (ls + rs ↝ e) +↝-return {e = e} ↝e = e , ↝e + +⟶-return : ∀ {A} {l : FST A} {e rs : List (FST A)} + → l + rs ⟶ e + → ∃[ e ] (l + rs ⟶ e) +⟶-return {e = e} ⟶e = e , ⟶e + +module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where + childs : FST A → List (FST A) + childs (pnode a as) = as mutual - infix 4 _+_⟶_ - data _+_⟶_ : PlainFST → List (PlainFST) → List (PlainFST) → Set where - base : ∀ {l : PlainFST} - --------------- - → l + [] ⟶ l ∷ [] - - merge : ∀ {a as bs rs cs} - → as + bs ↝ cs - ---------------------------------------------- - → pnode a as + pnode a bs ∷ rs ⟶ pnode a cs ∷ rs - - skip : ∀ {a as b bs rs cs} - → ¬ (a ≡ b) - → pnode a as + rs ⟶ cs - ---------------------------------------------- - → pnode a as + pnode b bs ∷ rs ⟶ pnode b bs ∷ cs - - -- This is bascially just a fold on lists. Maybe we can simplify it accordingly. - infix 4 _+_↝_ - data _+_↝_ : List PlainFST → List PlainFST → List PlainFST → Set where - impose-nothing : ∀ {rs} - → [] + rs ↝ rs - - impose-step : ∀ {l ls rs e e'} - → l + rs ⟶ e' - → ls + e' ↝ e - ---------------- - → l ∷ ls + rs ↝ e + ↝-total : ∀ (ls rs : List (FST A)) → ∃[ e ] (ls + rs ↝ e) + ↝-total [] rs = ↝-return impose-nothing + ↝-total (pnode a as ∷ ls) rs = + let e' , ⟶e' = ⟶-total (pnode a as) rs (↝-total as) + _ , ↝e = ↝-total ls e' + in ↝-return (impose-step ⟶e' ↝e) + + ⟶-total : ∀ (l : FST A) (rs : List (FST A)) → ((rs' : List (FST A)) → ∃[ e ] (childs l + rs' ↝ e)) → ∃[ e ] (l + rs ⟶ e) + ⟶-total l [] _ = ⟶-return base + ⟶-total (pnode a as) (pnode b bs ∷ rs) ↝-total-as with a ≟ b + ... | yes refl = + let cs , ↝cs = ↝-total-as bs + in ⟶-return (merge ↝cs) + ... | no a≠b = + let cs , ⟶cs = ⟶-total (pnode a as) rs ↝-total-as + in ⟶-return (skip a≠b ⟶cs) + + pdifferent : Rel (FST A) 0ℓ + pdifferent (pnode a _) (pnode b _) = False (a ≟ b) + + map-pdifferent : ∀ {b xs} (ys : List (FST A)) (z : FST A) + → pdifferent (pnode b xs) z + → pdifferent (pnode b ys) z + map-pdifferent {b} _ (pnode z _) l with z ≟ b + ... | yes _ = l + ... | no _ = l + + map-all-pdifferent : ∀ {b cs cs' xs} + → All (pdifferent (pnode b cs )) xs + → All (pdifferent (pnode b cs')) xs + map-all-pdifferent [] = [] + map-all-pdifferent {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-pdifferent cs' x px ∷ map-all-pdifferent pxs + + Unique : List (FST A) → Set + Unique = AllPairs pdifferent + + unique-ignores-children : ∀ {a as bs rs} + → Unique (pnode a as ∷ rs) + → Unique (pnode a bs ∷ rs) + unique-ignores-children (x ∷ xs) = map-all-pdifferent x ∷ xs + + drop-second-Unique : ∀ {x y zs} + → Unique (x ∷ y ∷ zs) + → Unique (x ∷ zs) + drop-second-Unique ((_ ∷ pxs) ∷ _ ∷ zs) = pxs ∷ zs mutual - ↝-deterministic : ∀ {fs rs e e'} - → fs + rs ↝ e - → fs + rs ↝ e' - → e ≡ e' - ↝-deterministic impose-nothing impose-nothing = refl - ↝-deterministic (impose-step ⟶x ↝x) (impose-step ⟶y ↝y) rewrite ⟶-deterministic ⟶x ⟶y | ↝-deterministic ↝x ↝y = refl - - ⟶-deterministic : ∀ {f rs e e'} - → f + rs ⟶ e - → f + rs ⟶ e' - → e ≡ e' - ⟶-deterministic base base = refl - ⟶-deterministic (merge x) (merge y) rewrite ↝-deterministic x y = refl - ⟶-deterministic (merge x) (skip a≠a y) = ⊥-elim (a≠a refl) - ⟶-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) - ⟶-deterministic (skip neq x) (skip neq' y) rewrite ⟶-deterministic x y = refl - - ↝-return : ∀ {e ls rs} - → ls + rs ↝ e - → ∃[ e ] (ls + rs ↝ e) - ↝-return {e} ↝e = e , ↝e - - ⟶-return : ∀ {e l rs} - → l + rs ⟶ e - → ∃[ e ] (l + rs ⟶ e) - ⟶-return {e} ⟶e = e , ⟶e - - module Impose (_≟_ : DecidableEquality A) where - - childs : PlainFST → List PlainFST - childs (pnode a as) = as - - mutual - ↝-total : ∀ (ls rs : List PlainFST) → ∃[ e ] (ls + rs ↝ e) - ↝-total [] rs = ↝-return impose-nothing - ↝-total (pnode a as ∷ ls) rs = - let e' , ⟶e' = ⟶-total (pnode a as) rs (↝-total as) - _ , ↝e = ↝-total ls e' - in ↝-return (impose-step ⟶e' ↝e) - - ⟶-total : ∀ (l : PlainFST) (rs : List PlainFST) → ((rs' : List PlainFST) → ∃[ e ] (childs l + rs' ↝ e)) → ∃[ e ] (l + rs ⟶ e) - ⟶-total l [] _ = ⟶-return base - ⟶-total (pnode a as) (pnode b bs ∷ rs) ↝-total-as with a ≟ b - ... | yes refl = - let cs , ↝cs = ↝-total-as bs - in ⟶-return (merge ↝cs) - ... | no a≠b = - let cs , ⟶cs = ⟶-total (pnode a as) rs ↝-total-as - in ⟶-return (skip a≠b ⟶cs) - - pdifferent : Rel PlainFST 0ℓ - pdifferent (pnode a _) (pnode b _) = False (a ≟ b) - - map-pdifferent : ∀ {b xs} (ys : List PlainFST) (z : PlainFST) - → pdifferent (pnode b xs) z - → pdifferent (pnode b ys) z - map-pdifferent {b} _ (pnode z _) l with z ≟ b - ... | yes _ = l - ... | no _ = l - - map-all-pdifferent : ∀ {b cs cs' xs} - → All (pdifferent (pnode b cs )) xs - → All (pdifferent (pnode b cs')) xs - map-all-pdifferent [] = [] - map-all-pdifferent {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-pdifferent cs' x px ∷ map-all-pdifferent pxs - - Unique : List PlainFST → Set - Unique = AllPairs pdifferent - - unique-ignores-children : ∀ {a as bs rs} - → Unique (pnode a as ∷ rs) - → Unique (pnode a bs ∷ rs) - unique-ignores-children (x ∷ xs) = map-all-pdifferent x ∷ xs - - drop-second-Unique : ∀ {x y zs} - → Unique (x ∷ y ∷ zs) - → Unique (x ∷ zs) - drop-second-Unique ((_ ∷ pxs) ∷ _ ∷ zs) = pxs ∷ zs - - mutual - data UniqueNode : PlainFST → Set where - unq : ∀ {a cs} → UniqueR cs → UniqueNode (pnode a cs) + data UniqueNode : FST A → Set where + unq : ∀ {a cs} → UniqueR cs → UniqueNode (pnode a cs) - UniqueR : List PlainFST → Set - UniqueR cs = Unique cs × All UniqueNode cs + UniqueR : List (FST A) → Set + UniqueR cs = Unique cs × All UniqueNode cs - mutual - ↝-preserves-unique : ∀ {ls rs e : List PlainFST} - → ls + rs ↝ e - → UniqueR ls - → UniqueR rs - → UniqueR e - ↝-preserves-unique impose-nothing ur-ls ur-rs = ur-rs - ↝-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ⟶e' ↝e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = - let ur-e' = ⟶-preserves-unique a as rs e' ⟶e' ur-as ur-rs - ur-e = ↝-preserves-unique ↝e (u-ls , ur-ls) ur-e' - in ur-e - - ⟶-preserves-unique : ∀ (a : A) (ls rs : List PlainFST) (e : List PlainFST) - → pnode a ls + rs ⟶ e - → UniqueR ls - → UniqueR rs - → UniqueR e - ⟶-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] - ⟶-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↝e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) - = unique-ignores-children u-rs , unq (↝-preserves-unique ↝e ur-ls ur-bs) ∷ un-rs - ⟶-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) - = induction a≠b (u-r ∷ u-rs) ⟶cs ∷ u-cs , ur-bs ∷ un-cs - where - ur-cs = ⟶-preserves-unique a ls rs cs ⟶cs u-ls (u-rs , ur-rs) - u-cs = proj₁ ur-cs - un-cs = proj₂ ur-cs - - induction : ∀ {a ls rs cs b bs} - → ¬ (a ≡ b) - → Unique (pnode b bs ∷ rs) - → pnode a ls + rs ⟶ cs - → All (pdifferent (pnode b bs)) cs - induction a≠b _ base = False-sym _≟_ (≠→False _≟_ a≠b) ∷ [] - induction a≠b u-rs (merge _) = False-sym _≟_ (≠→False _≟_ a≠b) ∷ head (drop-second-Unique u-rs) - induction a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ induction a≠b (u-r ∷ u-rs) ⟶cs - - ---- SPL Stuff ---- - - -- Feature Structure Forest - FSF : Set - FSF = Σ (List PlainFST) UniqueR - - forget-uniqueness : FSF → List PlainFST - forget-uniqueness = proj₁ - - infixr 3 _::_ - record Feature (N : 𝔽) : Set where - constructor _::_ - field - name : Name N - impl : FSF - open Feature public - - SPL : (N : 𝔽) → Set --𝔼 - SPL N = List (Feature N) - - select : ∀ {N} → Conf N → SPL N → SPL N - select c = filterᵇ (c ∘ name) - - forget-names : ∀ {N} → SPL N → List FSF - forget-names = map impl - - names : ∀ {N} → SPL N → List N - names = map name - - ---- Algebra ---- - open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) - open Eq.≡-Reasoning - - 𝟘 : FSF - 𝟘 = [] , [] , [] - - infixr 7 _⊕_ - _⊕_ : FSF → FSF → FSF - (l , u-l) ⊕ (r , u-r) = - let e , ↝e = ↝-total l r - u-e = ↝-preserves-unique ↝e u-l u-r - in e , u-e - - ⊕-all : List FSF → FSF - ⊕-all = foldr _⊕_ 𝟘 - - l-id : LeftIdentity _≡_ 𝟘 _⊕_ - l-id _ = refl - - r-id : RightIdentity _≡_ 𝟘 _⊕_ - r-id ([] , [] , []) = refl - r-id (.(pnode _ _) ∷ [] , [] ∷ [] , unq x ∷ []) = refl - r-id (x ∷ y ∷ zs , u-x ∷ u-y ∷ u-zs , ur-x ∷ ur-y ∷ ur-zs) = {!!} - - assoc : Associative _≡_ _⊕_ - assoc x y z = {!!} - - cong : Congruent₂ _≡_ _⊕_ - cong refl refl = refl - - idem : ∀ (i₁ i₂ : FSF) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ - idem = {!!} - - FST-is-FeatureAlgebra : FeatureAlgebra FSF _⊕_ 𝟘 - FST-is-FeatureAlgebra = record - { monoid = record - { isSemigroup = record - { isMagma = record - { isEquivalence = Eq.isEquivalence - ; ∙-cong = cong - } - ; assoc = assoc + mutual + ↝-preserves-unique : ∀ {ls rs e : List (FST A)} + → ls + rs ↝ e + → UniqueR ls + → UniqueR rs + → UniqueR e + ↝-preserves-unique impose-nothing ur-ls ur-rs = ur-rs + ↝-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ⟶e' ↝e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = + let ur-e' = ⟶-preserves-unique a as rs e' ⟶e' ur-as ur-rs + ur-e = ↝-preserves-unique ↝e (u-ls , ur-ls) ur-e' + in ur-e + + ⟶-preserves-unique : ∀ (a : A) (ls rs : List (FST A)) (e : List (FST A)) + → pnode a ls + rs ⟶ e + → UniqueR ls + → UniqueR rs + → UniqueR e + ⟶-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] + ⟶-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↝e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) + = unique-ignores-children u-rs , unq (↝-preserves-unique ↝e ur-ls ur-bs) ∷ un-rs + ⟶-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) + = ind a≠b (u-r ∷ u-rs) ⟶cs ∷ u-cs , ur-bs ∷ un-cs + where + ur-cs = ⟶-preserves-unique a ls rs cs ⟶cs u-ls (u-rs , ur-rs) + u-cs = proj₁ ur-cs + un-cs = proj₂ ur-cs + + ind : ∀ {a ls rs cs b bs} + → ¬ (a ≡ b) + → Unique (pnode b bs ∷ rs) + → pnode a ls + rs ⟶ cs + → All (pdifferent (pnode b bs)) cs + ind a≠b _ base = False-sym _≟_ (≠→False _≟_ a≠b) ∷ [] + ind a≠b u-rs (merge _) = False-sym _≟_ (≠→False _≟_ a≠b) ∷ head (drop-second-Unique u-rs) + ind a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ ind a≠b (u-r ∷ u-rs) ⟶cs + + -- Feature Structure Forest + record FSF : Set where + constructor _⊚_ + field + trees : List (FST A) + valid : UniqueR trees + open FSF public + + forget-uniqueness : FSF → List (FST A) + forget-uniqueness = trees + + toVariants : FSF → List (Rose ∞ A) + toVariants = map toVariant ∘ trees + + {- + A feature is a named feature structure tree. + All features in a product line are required to have + the very same root node, otherwise they could not be + imposed. + To ensure this constraint by design, this root node is + part of the SPL definition and not the feature. + Hence, a feature is a rootless tree: It holds a list of trees, + which denote the children of the common root. + -} + infixr 3 _::_ + record Feature : Set where + constructor _::_ + field + name : Name F + impl : FSF + open Feature public + + record SPL : Set where + constructor _◀_ + field + root : A + features : List Feature + open SPL public + + select : Conf → List Feature → List Feature + select c = filterᵇ (c ∘ name) + + -- forget-names : ∀ {N} → SPL N → List FSF + -- forget-names = map impl + + names : SPL → List F + names = map name ∘ features + + ---- Algebra ---- + open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) + open Eq.≡-Reasoning + + 𝟘 : FSF + 𝟘 = [] ⊚ ([] , []) + + infixr 7 _⊕_ + _⊕_ : FSF → FSF → FSF + (l ⊚ u-l) ⊕ (r ⊚ u-r) = + let e , ↝e = ↝-total l r + u-e = ↝-preserves-unique ↝e u-l u-r + in e ⊚ u-e + + ⊕-all : List FSF → FSF + ⊕-all = foldr _⊕_ 𝟘 + + ⟦_⟧ : SPL → Conf → Rose ∞ A + ⟦ r ◀ features ⟧ c = rose (r -< toVariants (⊕-all (map impl (select c features))) >-) + + l-id : LeftIdentity _≡_ 𝟘 _⊕_ + l-id _ = refl + + r-id : RightIdentity _≡_ 𝟘 _⊕_ + r-id = {!!} + -- r-id ([] ⊚ ([] , [])) = refl + -- r-id (.(pnode _ _) ∷ [] , [] ∷ [] , unq x ∷ []) = refl + -- r-id (x ∷ y ∷ zs , u-x ∷ u-y ∷ u-zs , ur-x ∷ ur-y ∷ ur-zs) = {!!} + + assoc : Associative _≡_ _⊕_ + assoc x y z = {!!} + + cong : Congruent₂ _≡_ _⊕_ + cong refl refl = refl + + idem : ∀ (i₁ i₂ : FSF) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ + idem = {!!} + + FST-is-FeatureAlgebra : FeatureAlgebra FSF _⊕_ 𝟘 + FST-is-FeatureAlgebra = record + { monoid = record + { isSemigroup = record + { isMagma = record + { isEquivalence = Eq.isEquivalence + ; ∙-cong = cong } - ; identity = l-id , r-id + ; assoc = assoc } - ; distant-idempotence = idem + ; identity = l-id , r-id } - where - open import Data.Product using (_,_) + ; distant-idempotence = idem + } + where + open import Data.Product using (_,_) - ⟦_⟧ : ∀ {N : 𝔽} → SPL N → Conf N → FSF - ⟦ features ⟧ c = (⊕-all ∘ forget-names ∘ select c) features + -- We could avoid wrap and unwrap by defining our own intermediate tree structure + -- that does not reuse Artifact constructor. + -- unwrap : Rose A → Artifact A (Rose A) + -- unwrap (artifact a) = a - -- We could avoid wrap and unwrap by defining our own intermediate tree structure - -- that does not reuse Artifact constructor. - -- unwrap : Rose A → Artifact A (Rose A) - -- unwrap (artifact a) = a + -- wrap : Artifact A (Rose A) → Rose A + -- wrap a = artifact a - -- wrap : Artifact A (Rose A) → Rose A - -- wrap a = artifact a + open import Data.String using (String; _<+>_) + open import Show.Lines - open import Data.String using (String; _<+>_) - open import Show.Lines - - module Show {N : 𝔽} (show-N : N → String) (show-A : A → String) where - mutual - show-FST : PlainFST → Lines - show-FST = PlainFST-induction λ a children → do - > show-A a - indent 2 (lines children) - - show-FSF : List PlainFST → Lines - show-FSF roots = lines (map show-FST roots) - - show-Feature : Feature N → Lines - show-Feature feature = do - > show-N (name feature) <+> "∷" - indent 2 (show-FSF (forget-uniqueness (impl feature))) + module Show (show-F : F → String) (show-A : A → String) where + mutual + show-FST : FST A → Lines + show-FST = induction λ a children → do + > show-A a + indent 2 (lines children) + + show-FSF : List (FST A) → Lines + show-FSF roots = lines (map show-FST roots) + + show-Feature : Feature → Lines + show-Feature feature = do + > show-F (name feature) <+> "∷" + indent 2 (show-FSF (forget-uniqueness (impl feature))) From 73cd3db5b29d0942d903bdfbc6c15e5e92edb174 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Tue, 19 Mar 2024 17:27:50 +0100 Subject: [PATCH 27/43] much cleaner relations --- src/Lang/FST.agda | 70 +++++++++++++++++++++++++++++------------------ 1 file changed, 43 insertions(+), 27 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 2c70e4d9..bad851c4 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -17,7 +17,7 @@ open import Size using (∞) open import Relation.Nullary.Negation using (¬_) open import Relation.Nullary.Decidable using (yes; no; _because_; False) -open import Relation.Binary using (DecidableEquality; Rel) +open import Relation.Binary using (Decidable; DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) open import Framework.Annotation.Name using (Name) @@ -65,6 +65,32 @@ toVariant {A} = induction step step : A → List (Rose ∞ A) → Rose ∞ A step a cs = rose (a -< cs >-) +_≈_ : ∀ {A} → Rel (FST A) 0ℓ +(pnode a _) ≈ (pnode b _) = a ≡ b + +≈-sym : ∀ {A} → (a b : FST A) → a ≈ b → b ≈ a +≈-sym (pnode a _) (pnode .a _) refl = refl + +_≉_ : ∀ {A} → Rel (FST A) 0ℓ +a ≉ b = ¬ (a ≈ b) + +≉-sym : ∀ {A} → (a b : FST A) → a ≉ b → b ≉ a +≉-sym a b a≉b b≈a = a≉b (≈-sym b a b≈a) + +_∉_ : ∀ {A} → FST A → List (FST A) → Set +x ∉ xs = All (x ≉_) xs + +map-≉ : ∀ {A} {b xs} (ys : List (FST A)) (z : FST A) + → pnode b xs ≉ z + → pnode b ys ≉ z +map-≉ ys (pnode z zs) z≉z refl = z≉z refl + +map-∉ : ∀ {A} {b : A} {cs cs' xs : List (FST A)} + → pnode b cs ∉ xs + → pnode b cs' ∉ xs +map-∉ [] = [] +map-∉ {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-≉ cs' x px ∷ map-∉ pxs + -- the syntax used in the paper for paths infixr 5 _._ _._ : ∀ {A} → A → (cs : List (FST A)) → List (FST A) @@ -86,11 +112,14 @@ mutual ---------------------------------------------- → pnode a as + pnode a bs ∷ rs ⟶ pnode a cs ∷ rs - skip : ∀ {A} {a b : A} {as bs rs cs : List (FST A)} - → ¬ (a ≡ b) - → pnode a as + rs ⟶ cs + -- In the original work, skipped nodes were added to the left. + -- We add to the right here because it fits nicer with list construction _∷_ + -- Otherwise, we would have to backtrack when we found no match in rs. + skip : ∀ {A} {a r : FST A} {rs cs : List (FST A)} + → a ≉ r + → a + rs ⟶ cs ---------------------------------------------- - → pnode a as + pnode b bs ∷ rs ⟶ pnode b bs ∷ cs + → a + r ∷ rs ⟶ r ∷ cs -- This is basically just a fold on lists. Maybe we can simplify it accordingly. infix 4 _+_↝_ @@ -133,6 +162,9 @@ mutual ⟶-return {e = e} ⟶e = e , ⟶e module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where + _==_ : Decidable (_≈_ {A}) + _==_ (pnode a _) (pnode b _) = a ≟ b + childs : FST A → List (FST A) childs (pnode a as) = as @@ -154,29 +186,13 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where let cs , ⟶cs = ⟶-total (pnode a as) rs ↝-total-as in ⟶-return (skip a≠b ⟶cs) - pdifferent : Rel (FST A) 0ℓ - pdifferent (pnode a _) (pnode b _) = False (a ≟ b) - - map-pdifferent : ∀ {b xs} (ys : List (FST A)) (z : FST A) - → pdifferent (pnode b xs) z - → pdifferent (pnode b ys) z - map-pdifferent {b} _ (pnode z _) l with z ≟ b - ... | yes _ = l - ... | no _ = l - - map-all-pdifferent : ∀ {b cs cs' xs} - → All (pdifferent (pnode b cs )) xs - → All (pdifferent (pnode b cs')) xs - map-all-pdifferent [] = [] - map-all-pdifferent {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-pdifferent cs' x px ∷ map-all-pdifferent pxs - Unique : List (FST A) → Set - Unique = AllPairs pdifferent + Unique = AllPairs _≉_ unique-ignores-children : ∀ {a as bs rs} → Unique (pnode a as ∷ rs) → Unique (pnode a bs ∷ rs) - unique-ignores-children (x ∷ xs) = map-all-pdifferent x ∷ xs + unique-ignores-children (x ∷ xs) = map-∉ x ∷ xs drop-second-Unique : ∀ {x y zs} → Unique (x ∷ y ∷ zs) @@ -217,13 +233,13 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where u-cs = proj₁ ur-cs un-cs = proj₂ ur-cs - ind : ∀ {a ls rs cs b bs} + ind : ∀ {a ls b bs cs rs} → ¬ (a ≡ b) → Unique (pnode b bs ∷ rs) → pnode a ls + rs ⟶ cs - → All (pdifferent (pnode b bs)) cs - ind a≠b _ base = False-sym _≟_ (≠→False _≟_ a≠b) ∷ [] - ind a≠b u-rs (merge _) = False-sym _≟_ (≠→False _≟_ a≠b) ∷ head (drop-second-Unique u-rs) + → All (_≉_ (pnode b bs)) cs + ind {a} {ls} {b} {bs} a≠b _ base = ≉-sym (pnode a ls) (pnode b bs) a≠b ∷ [] + ind {a} {_} {b} {bs} {pnode .a cs ∷ _} a≠b u-rs (merge _) = ≉-sym (pnode a cs) (pnode b bs) a≠b ∷ head (drop-second-Unique u-rs) ind a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ ind a≠b (u-r ∷ u-rs) ⟶cs -- Feature Structure Forest From 5eee17433e451d841b4416c1c92452f4330fb3fc Mon Sep 17 00:00:00 2001 From: pmbittner Date: Tue, 19 Mar 2024 19:13:42 +0100 Subject: [PATCH 28/43] proof of right identity of FSTs --- src/Lang/FST.agda | 165 ++++++++++++++++++++++++++++++++++++++-------- 1 file changed, 138 insertions(+), 27 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index bad851c4..fbfbb479 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -5,7 +5,9 @@ open import Framework.Definitions module Lang.FST (F : 𝔽) where open import Data.Bool using (Bool; true; false; if_then_else_) -open import Data.List using (List; []; _∷_; foldr; map; filterᵇ; concat; reverse) +open import Data.List using (List; []; _∷_; _∷ʳ_; _++_; foldr; map; filterᵇ; concat; reverse) +open import Data.List.Properties using (++-identityʳ) +open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_; head) open import Data.Product using (Σ; ∃-syntax; _×_; _,_; proj₁; proj₂) @@ -29,25 +31,10 @@ Conf : Set Conf = F → Bool module TODO-MOVE-TO-AUX-OR-USE-STL where - ≠-sym : ∀ {ℓ} {A : Set ℓ} (a b : A) - → ¬ (a ≡ b) - → ¬ (b ≡ a) - ≠-sym a b a≠b refl = a≠b refl - - ≠→False : ∀ {ℓ} {A : Set ℓ} {a b : A} - → (_≟_ : DecidableEquality A) - → ¬ (a ≡ b) - → False (a ≟ b) - ≠→False {a = a} {b = b} _≟_ a≠b with a ≟ b - ... | yes a≡b = ⊥-elim (a≠b a≡b) - ... | no _ = tt - - False-sym : ∀ {ℓ} {A : Set ℓ} {a b : A} - → (_≟_ : DecidableEquality A) - → False (a ≟ b) - → False (b ≟ a) - False-sym {a = a} {b = b} _≟_ _ with a ≟ b - ... | no ¬p = ≠→False _≟_ (≠-sym a b ¬p) + lem : ∀ {ℓ} {A : Set ℓ} (y : A) (ys xs : List A) + → (xs ++ y ∷ []) ++ ys ≡ (xs ++ (y ∷ ys)) + lem y ys [] = refl + lem y ys (x ∷ xs) = Eq.cong (x ∷_) (lem y ys xs) open TODO-MOVE-TO-AUX-OR-USE-STL data FST : 𝔼 where @@ -77,9 +64,39 @@ a ≉ b = ¬ (a ≈ b) ≉-sym : ∀ {A} → (a b : FST A) → a ≉ b → b ≉ a ≉-sym a b a≉b b≈a = a≉b (≈-sym b a b≈a) +_∈_ : ∀ {A} → FST A → List (FST A) → Set +x ∈ xs = Any (x ≈_) xs + _∉_ : ∀ {A} → FST A → List (FST A) → Set x ∉ xs = All (x ≉_) xs +_⊑_ : ∀ {A} → (xs ys : List (FST A)) → Set --\squb= +xs ⊑ ys = All (_∈ ys) xs + +_⋢_ : ∀ {A} → (xs ys : List (FST A)) → Set --\squb=n +xs ⋢ ys = Any (_∉ ys) xs + +Disjoint : ∀ {A} → (xs ys : List (FST A)) → Set --\squb=n +Disjoint xs ys = All (_∉ ys) xs + +-- identity of proofs +open import Axioms.Extensionality using (extensionality) +≉-deterministic : ∀ {A} (x y : FST A) + → (p₁ : x ≉ y) + → (p₂ : x ≉ y) + → p₁ ≡ p₂ +≉-deterministic (pnode a _) (pnode b _) p₁ p₂ = extensionality λ where refl → refl + +∉-deterministic : ∀ {A} {x : FST A} (ys : List (FST A)) + → (p₁ : x ∉ ys) + → (p₂ : x ∉ ys) + → p₁ ≡ p₂ +∉-deterministic [] [] [] = refl +∉-deterministic {_} {x} (y ∷ ys) (x≉y₁ ∷ pa) (x≉y₂ ∷ pb) + rewrite ≉-deterministic x y x≉y₁ x≉y₂ + rewrite ∉-deterministic ys pa pb + = refl + map-≉ : ∀ {A} {b xs} (ys : List (FST A)) (z : FST A) → pnode b xs ≉ z → pnode b ys ≉ z @@ -91,6 +108,35 @@ map-∉ : ∀ {A} {b : A} {cs cs' xs : List (FST A)} map-∉ [] = [] map-∉ {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-≉ cs' x px ∷ map-∉ pxs +disjoint-[]ˡ : ∀ {A} (xs : List (FST A)) → Disjoint [] xs +disjoint-[]ˡ _ = [] + +disjoint-[]ʳ : ∀ {A} (xs : List (FST A)) → Disjoint xs [] +disjoint-[]ʳ [] = [] +disjoint-[]ʳ (x ∷ xs) = [] ∷ (disjoint-[]ʳ xs) + +disjoint-grow : ∀ {A} (r : FST A) (rs ls : List (FST A)) + → Disjoint ls rs + → r ∉ ls + → Disjoint ls (r ∷ rs) +disjoint-grow r rs [] _ _ = [] +disjoint-grow r rs (l ∷ ls) (l∉rs ∷ d-ls-rs) (r≉l ∷ r∉ls) + = (≉-sym r l r≉l ∷ l∉rs) ∷ disjoint-grow r rs ls d-ls-rs r∉ls + +disjoint-shiftʳ : ∀ {A} (r : FST A) (rs ls : List (FST A)) + → Disjoint ls (r ∷ rs) + → Disjoint ls (rs ++ r ∷ []) +disjoint-shiftʳ r rs [] x = [] +disjoint-shiftʳ r rs (l ∷ ls) ((l≉r ∷ l∉rs) ∷ d-ls-rrs) + = step l r rs l≉r l∉rs ∷ disjoint-shiftʳ r rs ls d-ls-rrs + where + step : ∀ {A} (x y : FST A) (zs : List (FST A)) + → x ≉ y + → x ∉ zs + → x ∉ (zs ++ y ∷ []) + step x y [] x≉y _ = x≉y ∷ [] + step x y (z ∷ zs) x≉y (x≉z ∷ x∉zs) = x≉z ∷ step x y zs x≉y x∉zs + -- the syntax used in the paper for paths infixr 5 _._ _._ : ∀ {A} → A → (cs : List (FST A)) → List (FST A) @@ -206,6 +252,27 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where UniqueR : List (FST A) → Set UniqueR cs = Unique cs × All UniqueNode cs + mutual + UniqueNode-deterministic : ∀ {x : FST A} + → (a : UniqueNode x) + → (b : UniqueNode x) + → a ≡ b + UniqueNode-deterministic {pnode _ cs} (unq a) (unq b) = Eq.cong unq (UniqueR-deterministic cs a b) + + UniqueR-deterministic : ∀ (xs : List (FST A)) + → (ua : UniqueR xs) + → (ub : UniqueR xs) + → ua ≡ ub + UniqueR-deterministic [] ([] , []) ([] , []) = refl + UniqueR-deterministic (x ∷ xs) (a-x∉xs ∷ a-u-xs , a-ur-x ∷ a-ur-xs) (b-x∉xs ∷ b-u-xs , b-ur-x ∷ b-ur-xs) + with UniqueR-deterministic xs (a-u-xs , a-ur-xs) (b-u-xs , b-ur-xs) + ... | eq + rewrite (Eq.cong proj₁ eq) + rewrite (Eq.cong proj₂ eq) + rewrite UniqueNode-deterministic a-ur-x b-ur-x + rewrite ∉-deterministic xs a-x∉xs b-x∉xs + = refl + mutual ↝-preserves-unique : ∀ {ls rs e : List (FST A)} → ls + rs ↝ e @@ -250,6 +317,30 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where valid : UniqueR trees open FSF public + ⟶-append-strange : ∀ (l : FST A) (rs : List (FST A)) + → l ∉ rs + → l + rs ⟶ rs ∷ʳ l + ⟶-append-strange l [] _ = base + ⟶-append-strange l (r ∷ rs) (l≠r ∷ l∉rs) = skip l≠r (⟶-append-strange l rs l∉rs) + + ↝-append-strangers : ∀ (ls rs : List (FST A)) + → Unique ls + → Disjoint ls rs + → ls + rs ↝ rs ++ ls + ↝-append-strangers [] rs _ _ rewrite ++-identityʳ rs = impose-nothing + ↝-append-strangers (l ∷ ls) rs (l∉ls ∷ u-ls) (l∉rs ∷ d-ls-rs) + rewrite (Eq.sym (lem l ls rs)) + with ⟶-append-strange l rs l∉rs + ... | x + = impose-step x (↝-append-strangers ls (rs ++ l ∷ []) u-ls + (disjoint-shiftʳ l rs ls (disjoint-grow l rs ls d-ls-rs l∉ls))) + + impose-nothing-r : + ∀ (ls : List (FST A)) + → Unique ls + → ls + [] ↝ ls + impose-nothing-r ls u-ls = ↝-append-strangers ls [] u-ls (disjoint-[]ʳ ls) + forget-uniqueness : FSF → List (FST A) forget-uniqueness = trees @@ -307,17 +398,33 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where ⊕-all : List FSF → FSF ⊕-all = foldr _⊕_ 𝟘 - ⟦_⟧ : SPL → Conf → Rose ∞ A - ⟦ r ◀ features ⟧ c = rose (r -< toVariants (⊕-all (map impl (select c features))) >-) - l-id : LeftIdentity _≡_ 𝟘 _⊕_ l-id _ = refl r-id : RightIdentity _≡_ 𝟘 _⊕_ - r-id = {!!} - -- r-id ([] ⊚ ([] , [])) = refl - -- r-id (.(pnode _ _) ∷ [] , [] ∷ [] , unq x ∷ []) = refl - -- r-id (x ∷ y ∷ zs , u-x ∷ u-y ∷ u-zs , ur-x ∷ ur-y ∷ ur-zs) = {!!} + r-id (xs ⊚ (u-xs , ur-xs)) + -- Let's see what ⊕ does + with ↝-total xs [] + -- it computes some result 'e' and a derivation 'deriv' + ... | (e , deriv) + -- However, we know by impose-nothing-r that we can derive + -- 'xs' itself as result. + -- By determinism, we know that there can only be one derivation + -- so e = xs. + -- (We can't do a rewrite here for some reason so we stick to good old "with".) + with ↝-deterministic deriv (impose-nothing-r xs u-xs) + ... | refl = Eq.cong (xs ⊚_) (help xs (u-xs , ur-xs) deriv) + where + -- lastly, we have to prove that the typing is also unique but that is actually + -- irrelevant. Maybe we can avoid this proof somehow? + -- Its never needed and just an artifical problem. + -- Maybe we shouldnt prove for _≡_ but rather for a new eq relation + -- that is weaker and ignores the typing. + help : ∀ (ls : List (FST A)) + → (ur-ls : UniqueR ls) + → (deriv : ls + [] ↝ ls) + → ↝-preserves-unique deriv ur-ls ([] , []) ≡ ur-ls + help ls ur-ls deriv = UniqueR-deterministic ls (↝-preserves-unique deriv ur-ls ([] , [])) ur-ls assoc : Associative _≡_ _⊕_ assoc x y z = {!!} @@ -345,6 +452,10 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where where open import Data.Product using (_,_) + -- Semantics + ⟦_⟧ : SPL → Conf → Rose ∞ A + ⟦ r ◀ features ⟧ c = rose (r -< toVariants (⊕-all (map impl (select c features))) >-) + -- We could avoid wrap and unwrap by defining our own intermediate tree structure -- that does not reuse Artifact constructor. -- unwrap : Rose A → Artifact A (Rose A) From 01f190f74a9598061529f7a2b247a5d6d1e3f2b1 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 20 Mar 2024 13:14:18 +0100 Subject: [PATCH 29/43] some renamings according to the paper --- src/Lang/FST.agda | 168 +++++++++++++++++++++++----------------------- 1 file changed, 84 insertions(+), 84 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index fbfbb479..1e7c2163 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -147,66 +147,66 @@ branches : ∀ {A} → List (List (FST A)) → List (FST A) branches = concat mutual - infix 4 _+_⟶_ - data _+_⟶_ : ∀ {A} → FST A → List (FST A) → List (FST A) → Set where + infix 4 _⊙_↝_ + data _⊙_↝_ : ∀ {A} → FST A → List (FST A) → List (FST A) → Set where base : ∀ {A} {l : FST A} --------------- - → l + [] ⟶ l ∷ [] + → l ⊙ [] ↝ l ∷ [] merge : ∀ {A} {a : A} {as bs rs cs : List (FST A)} - → as + bs ↝ cs + → as + bs ↪ cs ---------------------------------------------- - → pnode a as + pnode a bs ∷ rs ⟶ pnode a cs ∷ rs + → pnode a as ⊙ pnode a bs ∷ rs ↝ pnode a cs ∷ rs -- In the original work, skipped nodes were added to the left. -- We add to the right here because it fits nicer with list construction _∷_ -- Otherwise, we would have to backtrack when we found no match in rs. skip : ∀ {A} {a r : FST A} {rs cs : List (FST A)} → a ≉ r - → a + rs ⟶ cs + → a ⊙ rs ↝ cs ---------------------------------------------- - → a + r ∷ rs ⟶ r ∷ cs + → a ⊙ r ∷ rs ↝ r ∷ cs -- This is basically just a fold on lists. Maybe we can simplify it accordingly. - infix 4 _+_↝_ - data _+_↝_ : ∀ {A} → List (FST A) → List (FST A) → List (FST A) → Set where + infix 4 _+_↪_ + data _+_↪_ : ∀ {A} → List (FST A) → List (FST A) → List (FST A) → Set where impose-nothing : ∀ {A} {rs : List (FST A)} - → [] + rs ↝ rs + → [] + rs ↪ rs impose-step : ∀ {A} {l : FST A} {ls rs e e' : List (FST A)} - → l + rs ⟶ e' - → ls + e' ↝ e + → l ⊙ rs ↝ e' + → ls + e' ↪ e ---------------- - → l ∷ ls + rs ↝ e + → l ∷ ls + rs ↪ e mutual - ↝-deterministic : ∀ {A} {fs rs e e' : List (FST A)} - → fs + rs ↝ e - → fs + rs ↝ e' + ↪-deterministic : ∀ {A} {fs rs e e' : List (FST A)} + → fs + rs ↪ e + → fs + rs ↪ e' → e ≡ e' - ↝-deterministic impose-nothing impose-nothing = refl - ↝-deterministic (impose-step ⟶x ↝x) (impose-step ⟶y ↝y) rewrite ⟶-deterministic ⟶x ⟶y | ↝-deterministic ↝x ↝y = refl + ↪-deterministic impose-nothing impose-nothing = refl + ↪-deterministic (impose-step ↝x ↪x) (impose-step ↝y ↪y) rewrite ↝-deterministic ↝x ↝y | ↪-deterministic ↪x ↪y = refl - ⟶-deterministic : ∀ {A} {f : FST A} {rs e e' : List (FST A)} - → f + rs ⟶ e - → f + rs ⟶ e' + ↝-deterministic : ∀ {A} {f : FST A} {rs e e' : List (FST A)} + → f ⊙ rs ↝ e + → f ⊙ rs ↝ e' → e ≡ e' - ⟶-deterministic base base = refl - ⟶-deterministic (merge x) (merge y) rewrite ↝-deterministic x y = refl - ⟶-deterministic (merge x) (skip a≠a y) = ⊥-elim (a≠a refl) - ⟶-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) - ⟶-deterministic (skip neq x) (skip neq' y) rewrite ⟶-deterministic x y = refl - -↝-return : ∀ {A} {e ls rs : List (FST A)} - → ls + rs ↝ e - → ∃[ e ] (ls + rs ↝ e) + ↝-deterministic base base = refl + ↝-deterministic (merge x) (merge y) rewrite ↪-deterministic x y = refl + ↝-deterministic (merge x) (skip a≠a y) = ⊥-elim (a≠a refl) + ↝-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) + ↝-deterministic (skip neq x) (skip neq' y) rewrite ↝-deterministic x y = refl + +↪-return : ∀ {A} {e ls rs : List (FST A)} + → ls + rs ↪ e + → ∃[ e ] (ls + rs ↪ e) +↪-return {e = e} ↪e = e , ↪e + +↝-return : ∀ {A} {l : FST A} {e rs : List (FST A)} + → l ⊙ rs ↝ e + → ∃[ e ] (l ⊙ rs ↝ e) ↝-return {e = e} ↝e = e , ↝e -⟶-return : ∀ {A} {l : FST A} {e rs : List (FST A)} - → l + rs ⟶ e - → ∃[ e ] (l + rs ⟶ e) -⟶-return {e = e} ⟶e = e , ⟶e - module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where _==_ : Decidable (_≈_ {A}) _==_ (pnode a _) (pnode b _) = a ≟ b @@ -215,22 +215,22 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where childs (pnode a as) = as mutual - ↝-total : ∀ (ls rs : List (FST A)) → ∃[ e ] (ls + rs ↝ e) - ↝-total [] rs = ↝-return impose-nothing - ↝-total (pnode a as ∷ ls) rs = - let e' , ⟶e' = ⟶-total (pnode a as) rs (↝-total as) - _ , ↝e = ↝-total ls e' - in ↝-return (impose-step ⟶e' ↝e) - - ⟶-total : ∀ (l : FST A) (rs : List (FST A)) → ((rs' : List (FST A)) → ∃[ e ] (childs l + rs' ↝ e)) → ∃[ e ] (l + rs ⟶ e) - ⟶-total l [] _ = ⟶-return base - ⟶-total (pnode a as) (pnode b bs ∷ rs) ↝-total-as with a ≟ b + ↪-total : ∀ (ls rs : List (FST A)) → ∃[ e ] (ls + rs ↪ e) + ↪-total [] rs = ↪-return impose-nothing + ↪-total (pnode a as ∷ ls) rs = + let e' , ↝e' = ↝-total (pnode a as) rs (↪-total as) + _ , ↪e = ↪-total ls e' + in ↪-return (impose-step ↝e' ↪e) + + ↝-total : ∀ (l : FST A) (rs : List (FST A)) → ((rs' : List (FST A)) → ∃[ e ] (childs l + rs' ↪ e)) → ∃[ e ] (l ⊙ rs ↝ e) + ↝-total l [] _ = ↝-return base + ↝-total (pnode a as) (pnode b bs ∷ rs) ↪-total-as with a ≟ b ... | yes refl = - let cs , ↝cs = ↝-total-as bs - in ⟶-return (merge ↝cs) + let cs , ↪cs = ↪-total-as bs + in ↝-return (merge ↪cs) ... | no a≠b = - let cs , ⟶cs = ⟶-total (pnode a as) rs ↝-total-as - in ⟶-return (skip a≠b ⟶cs) + let cs , ↝cs = ↝-total (pnode a as) rs ↪-total-as + in ↝-return (skip a≠b ↝cs) Unique : List (FST A) → Set Unique = AllPairs _≉_ @@ -274,40 +274,40 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where = refl mutual - ↝-preserves-unique : ∀ {ls rs e : List (FST A)} - → ls + rs ↝ e + ↪-preserves-unique : ∀ {ls rs e : List (FST A)} + → ls + rs ↪ e → UniqueR ls → UniqueR rs → UniqueR e - ↝-preserves-unique impose-nothing ur-ls ur-rs = ur-rs - ↝-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ⟶e' ↝e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = - let ur-e' = ⟶-preserves-unique a as rs e' ⟶e' ur-as ur-rs - ur-e = ↝-preserves-unique ↝e (u-ls , ur-ls) ur-e' + ↪-preserves-unique impose-nothing ur-ls ur-rs = ur-rs + ↪-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ↝e' ↪e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = + let ur-e' = ↝-preserves-unique a as rs e' ↝e' ur-as ur-rs + ur-e = ↪-preserves-unique ↪e (u-ls , ur-ls) ur-e' in ur-e - ⟶-preserves-unique : ∀ (a : A) (ls rs : List (FST A)) (e : List (FST A)) - → pnode a ls + rs ⟶ e + ↝-preserves-unique : ∀ (a : A) (ls rs : List (FST A)) (e : List (FST A)) + → pnode a ls ⊙ rs ↝ e → UniqueR ls → UniqueR rs → UniqueR e - ⟶-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] - ⟶-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↝e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) - = unique-ignores-children u-rs , unq (↝-preserves-unique ↝e ur-ls ur-bs) ∷ un-rs - ⟶-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ⟶cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) - = ind a≠b (u-r ∷ u-rs) ⟶cs ∷ u-cs , ur-bs ∷ un-cs + ↝-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] + ↝-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↪e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) + = unique-ignores-children u-rs , unq (↪-preserves-unique ↪e ur-ls ur-bs) ∷ un-rs + ↝-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ↝cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) + = ind a≠b (u-r ∷ u-rs) ↝cs ∷ u-cs , ur-bs ∷ un-cs where - ur-cs = ⟶-preserves-unique a ls rs cs ⟶cs u-ls (u-rs , ur-rs) + ur-cs = ↝-preserves-unique a ls rs cs ↝cs u-ls (u-rs , ur-rs) u-cs = proj₁ ur-cs un-cs = proj₂ ur-cs ind : ∀ {a ls b bs cs rs} → ¬ (a ≡ b) → Unique (pnode b bs ∷ rs) - → pnode a ls + rs ⟶ cs + → pnode a ls ⊙ rs ↝ cs → All (_≉_ (pnode b bs)) cs ind {a} {ls} {b} {bs} a≠b _ base = ≉-sym (pnode a ls) (pnode b bs) a≠b ∷ [] ind {a} {_} {b} {bs} {pnode .a cs ∷ _} a≠b u-rs (merge _) = ≉-sym (pnode a cs) (pnode b bs) a≠b ∷ head (drop-second-Unique u-rs) - ind a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ⟶cs) = b≠b' ∷ ind a≠b (u-r ∷ u-rs) ⟶cs + ind a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ↝cs) = b≠b' ∷ ind a≠b (u-r ∷ u-rs) ↝cs -- Feature Structure Forest record FSF : Set where @@ -317,29 +317,29 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where valid : UniqueR trees open FSF public - ⟶-append-strange : ∀ (l : FST A) (rs : List (FST A)) + ↝-append-strange : ∀ (l : FST A) (rs : List (FST A)) → l ∉ rs - → l + rs ⟶ rs ∷ʳ l - ⟶-append-strange l [] _ = base - ⟶-append-strange l (r ∷ rs) (l≠r ∷ l∉rs) = skip l≠r (⟶-append-strange l rs l∉rs) + → l ⊙ rs ↝ rs ∷ʳ l + ↝-append-strange l [] _ = base + ↝-append-strange l (r ∷ rs) (l≠r ∷ l∉rs) = skip l≠r (↝-append-strange l rs l∉rs) - ↝-append-strangers : ∀ (ls rs : List (FST A)) + ↪-append-strangers : ∀ (ls rs : List (FST A)) → Unique ls → Disjoint ls rs - → ls + rs ↝ rs ++ ls - ↝-append-strangers [] rs _ _ rewrite ++-identityʳ rs = impose-nothing - ↝-append-strangers (l ∷ ls) rs (l∉ls ∷ u-ls) (l∉rs ∷ d-ls-rs) + → ls + rs ↪ rs ++ ls + ↪-append-strangers [] rs _ _ rewrite ++-identityʳ rs = impose-nothing + ↪-append-strangers (l ∷ ls) rs (l∉ls ∷ u-ls) (l∉rs ∷ d-ls-rs) rewrite (Eq.sym (lem l ls rs)) - with ⟶-append-strange l rs l∉rs + with ↝-append-strange l rs l∉rs ... | x - = impose-step x (↝-append-strangers ls (rs ++ l ∷ []) u-ls + = impose-step x (↪-append-strangers ls (rs ++ l ∷ []) u-ls (disjoint-shiftʳ l rs ls (disjoint-grow l rs ls d-ls-rs l∉ls))) impose-nothing-r : ∀ (ls : List (FST A)) → Unique ls - → ls + [] ↝ ls - impose-nothing-r ls u-ls = ↝-append-strangers ls [] u-ls (disjoint-[]ʳ ls) + → ls + [] ↪ ls + impose-nothing-r ls u-ls = ↪-append-strangers ls [] u-ls (disjoint-[]ʳ ls) forget-uniqueness : FSF → List (FST A) forget-uniqueness = trees @@ -391,8 +391,8 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where infixr 7 _⊕_ _⊕_ : FSF → FSF → FSF (l ⊚ u-l) ⊕ (r ⊚ u-r) = - let e , ↝e = ↝-total l r - u-e = ↝-preserves-unique ↝e u-l u-r + let e , ↪e = ↪-total l r + u-e = ↪-preserves-unique ↪e u-l u-r in e ⊚ u-e ⊕-all : List FSF → FSF @@ -404,7 +404,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where r-id : RightIdentity _≡_ 𝟘 _⊕_ r-id (xs ⊚ (u-xs , ur-xs)) -- Let's see what ⊕ does - with ↝-total xs [] + with ↪-total xs [] -- it computes some result 'e' and a derivation 'deriv' ... | (e , deriv) -- However, we know by impose-nothing-r that we can derive @@ -412,7 +412,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- By determinism, we know that there can only be one derivation -- so e = xs. -- (We can't do a rewrite here for some reason so we stick to good old "with".) - with ↝-deterministic deriv (impose-nothing-r xs u-xs) + with ↪-deterministic deriv (impose-nothing-r xs u-xs) ... | refl = Eq.cong (xs ⊚_) (help xs (u-xs , ur-xs) deriv) where -- lastly, we have to prove that the typing is also unique but that is actually @@ -422,9 +422,9 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- that is weaker and ignores the typing. help : ∀ (ls : List (FST A)) → (ur-ls : UniqueR ls) - → (deriv : ls + [] ↝ ls) - → ↝-preserves-unique deriv ur-ls ([] , []) ≡ ur-ls - help ls ur-ls deriv = UniqueR-deterministic ls (↝-preserves-unique deriv ur-ls ([] , [])) ur-ls + → (deriv : ls + [] ↪ ls) + → ↪-preserves-unique deriv ur-ls ([] , []) ≡ ur-ls + help ls ur-ls deriv = UniqueR-deterministic ls (↪-preserves-unique deriv ur-ls ([] , [])) ur-ls assoc : Associative _≡_ _⊕_ assoc x y z = {!!} From ecae7058261793e0781aea40d3561b39472e6922 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 20 Mar 2024 15:11:54 +0100 Subject: [PATCH 30/43] FSTs are rose variants by definition --- src/Lang/FST.agda | 21 ++++++++------------- 1 file changed, 8 insertions(+), 13 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 1e7c2163..4df205e6 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -37,8 +37,10 @@ module TODO-MOVE-TO-AUX-OR-USE-STL where lem y ys (x ∷ xs) = Eq.cong (x ∷_) (lem y ys xs) open TODO-MOVE-TO-AUX-OR-USE-STL -data FST : 𝔼 where - pnode : ∀ {A : 𝔸} → A → List (FST A) → FST A +FST : 𝔼 +FST = Rose ∞ + +pattern pnode a cs = rose (a -< cs >-) induction : ∀ {A : 𝔸} {B : Set} → (A → List B → B) → FST A → B induction {A} {B} f n = go n [] where @@ -46,12 +48,6 @@ induction {A} {B} f n = go n [] where go (pnode a []) bs = f a (reverse bs) go (pnode a (c ∷ cs)) bs = go (pnode a cs) (go c [] ∷ bs) -toVariant : ∀ {A} → FST A → Rose ∞ A -toVariant {A} = induction step - module toVariant-implementation where - step : A → List (Rose ∞ A) → Rose ∞ A - step a cs = rose (a -< cs >-) - _≈_ : ∀ {A} → Rel (FST A) 0ℓ (pnode a _) ≈ (pnode b _) = a ≡ b @@ -222,7 +218,9 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where _ , ↪e = ↪-total ls e' in ↪-return (impose-step ↝e' ↪e) - ↝-total : ∀ (l : FST A) (rs : List (FST A)) → ((rs' : List (FST A)) → ∃[ e ] (childs l + rs' ↪ e)) → ∃[ e ] (l ⊙ rs ↝ e) + ↝-total : ∀ (l : FST A) (rs : List (FST A)) + → ((rs' : List (FST A)) → ∃[ e ] (childs l + rs' ↪ e)) + → ∃[ e ] (l ⊙ rs ↝ e) ↝-total l [] _ = ↝-return base ↝-total (pnode a as) (pnode b bs ∷ rs) ↪-total-as with a ≟ b ... | yes refl = @@ -344,9 +342,6 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where forget-uniqueness : FSF → List (FST A) forget-uniqueness = trees - toVariants : FSF → List (Rose ∞ A) - toVariants = map toVariant ∘ trees - {- A feature is a named feature structure tree. All features in a product line are required to have @@ -454,7 +449,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- Semantics ⟦_⟧ : SPL → Conf → Rose ∞ A - ⟦ r ◀ features ⟧ c = rose (r -< toVariants (⊕-all (map impl (select c features))) >-) + ⟦ r ◀ features ⟧ c = rose (r -< forget-uniqueness (⊕-all (map impl (select c features))) >-) -- We could avoid wrap and unwrap by defining our own intermediate tree structure -- that does not reuse Artifact constructor. From 25d67c50c7dcb9bf93f4d228daa7d16d76c45c33 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 20 Mar 2024 15:32:00 +0100 Subject: [PATCH 31/43] use common artifact syntax for FSTs --- src/Lang/FST.agda | 80 ++++++++++++++++++++++------------------------- 1 file changed, 38 insertions(+), 42 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 4df205e6..6048d800 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -25,7 +25,7 @@ open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) open import Framework.Annotation.Name using (Name) open import Framework.Variants using (Rose; rose) open import Framework.Composition.FeatureAlgebra -open import Construct.Artifact +open import Construct.Artifact as At using () Conf : Set Conf = F → Bool @@ -40,29 +40,33 @@ open TODO-MOVE-TO-AUX-OR-USE-STL FST : 𝔼 FST = Rose ∞ -pattern pnode a cs = rose (a -< cs >-) +pattern _-<_>- a cs = rose (a At.-< cs >-) induction : ∀ {A : 𝔸} {B : Set} → (A → List B → B) → FST A → B induction {A} {B} f n = go n [] where go : FST A → List B → B - go (pnode a []) bs = f a (reverse bs) - go (pnode a (c ∷ cs)) bs = go (pnode a cs) (go c [] ∷ bs) + go (a -< [] >-) bs = f a (reverse bs) + go (a -< c ∷ cs >-) bs = go (a -< cs >-) (go c [] ∷ bs) +infix 15 _≈_ _≈_ : ∀ {A} → Rel (FST A) 0ℓ -(pnode a _) ≈ (pnode b _) = a ≡ b +(a -< _ >-) ≈ (b -< _ >-) = a ≡ b ≈-sym : ∀ {A} → (a b : FST A) → a ≈ b → b ≈ a -≈-sym (pnode a _) (pnode .a _) refl = refl +≈-sym (a -< _ >-) (.a -< _ >-) refl = refl +infix 15 _≉_ _≉_ : ∀ {A} → Rel (FST A) 0ℓ a ≉ b = ¬ (a ≈ b) ≉-sym : ∀ {A} → (a b : FST A) → a ≉ b → b ≉ a ≉-sym a b a≉b b≈a = a≉b (≈-sym b a b≈a) +infix 15 _∈_ _∈_ : ∀ {A} → FST A → List (FST A) → Set x ∈ xs = Any (x ≈_) xs +infix 15 _∉_ _∉_ : ∀ {A} → FST A → List (FST A) → Set x ∉ xs = All (x ≉_) xs @@ -81,7 +85,7 @@ open import Axioms.Extensionality using (extensionality) → (p₁ : x ≉ y) → (p₂ : x ≉ y) → p₁ ≡ p₂ -≉-deterministic (pnode a _) (pnode b _) p₁ p₂ = extensionality λ where refl → refl +≉-deterministic (a -< _ >-) (b -< _ >-) p₁ p₂ = extensionality λ where refl → refl ∉-deterministic : ∀ {A} {x : FST A} (ys : List (FST A)) → (p₁ : x ∉ ys) @@ -94,13 +98,13 @@ open import Axioms.Extensionality using (extensionality) = refl map-≉ : ∀ {A} {b xs} (ys : List (FST A)) (z : FST A) - → pnode b xs ≉ z - → pnode b ys ≉ z -map-≉ ys (pnode z zs) z≉z refl = z≉z refl + → b -< xs >- ≉ z + → b -< ys >- ≉ z +map-≉ ys (z -< zs >-) z≉z refl = z≉z refl map-∉ : ∀ {A} {b : A} {cs cs' xs : List (FST A)} - → pnode b cs ∉ xs - → pnode b cs' ∉ xs + → b -< cs >- ∉ xs + → b -< cs' >- ∉ xs map-∉ [] = [] map-∉ {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-≉ cs' x px ∷ map-∉ pxs @@ -136,7 +140,7 @@ disjoint-shiftʳ r rs (l ∷ ls) ((l≉r ∷ l∉rs) ∷ d-ls-rrs) -- the syntax used in the paper for paths infixr 5 _._ _._ : ∀ {A} → A → (cs : List (FST A)) → List (FST A) -a . cs = pnode a cs ∷ [] +a . cs = a -< cs >- ∷ [] -- helper function when branching in paths branches : ∀ {A} → List (List (FST A)) → List (FST A) @@ -152,7 +156,7 @@ mutual merge : ∀ {A} {a : A} {as bs rs cs : List (FST A)} → as + bs ↪ cs ---------------------------------------------- - → pnode a as ⊙ pnode a bs ∷ rs ↝ pnode a cs ∷ rs + → a -< as >- ⊙ a -< bs >- ∷ rs ↝ a -< cs >- ∷ rs -- In the original work, skipped nodes were added to the left. -- We add to the right here because it fits nicer with list construction _∷_ @@ -205,16 +209,16 @@ mutual module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where _==_ : Decidable (_≈_ {A}) - _==_ (pnode a _) (pnode b _) = a ≟ b + (a -< _ >-) == (b -< _ >-) = a ≟ b childs : FST A → List (FST A) - childs (pnode a as) = as + childs (a -< as >-) = as mutual ↪-total : ∀ (ls rs : List (FST A)) → ∃[ e ] (ls + rs ↪ e) ↪-total [] rs = ↪-return impose-nothing - ↪-total (pnode a as ∷ ls) rs = - let e' , ↝e' = ↝-total (pnode a as) rs (↪-total as) + ↪-total (a -< as >- ∷ ls) rs = + let e' , ↝e' = ↝-total (a -< as >-) rs (↪-total as) _ , ↪e = ↪-total ls e' in ↪-return (impose-step ↝e' ↪e) @@ -222,20 +226,20 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where → ((rs' : List (FST A)) → ∃[ e ] (childs l + rs' ↪ e)) → ∃[ e ] (l ⊙ rs ↝ e) ↝-total l [] _ = ↝-return base - ↝-total (pnode a as) (pnode b bs ∷ rs) ↪-total-as with a ≟ b + ↝-total (a -< as >-) (b -< bs >- ∷ rs) ↪-total-as with a ≟ b ... | yes refl = let cs , ↪cs = ↪-total-as bs in ↝-return (merge ↪cs) ... | no a≠b = - let cs , ↝cs = ↝-total (pnode a as) rs ↪-total-as + let cs , ↝cs = ↝-total (a -< as >-) rs ↪-total-as in ↝-return (skip a≠b ↝cs) Unique : List (FST A) → Set Unique = AllPairs _≉_ unique-ignores-children : ∀ {a as bs rs} - → Unique (pnode a as ∷ rs) - → Unique (pnode a bs ∷ rs) + → Unique (a -< as >- ∷ rs) + → Unique (a -< bs >- ∷ rs) unique-ignores-children (x ∷ xs) = map-∉ x ∷ xs drop-second-Unique : ∀ {x y zs} @@ -245,7 +249,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where mutual data UniqueNode : FST A → Set where - unq : ∀ {a cs} → UniqueR cs → UniqueNode (pnode a cs) + unq : ∀ {a cs} → UniqueR cs → UniqueNode (a -< cs >-) UniqueR : List (FST A) → Set UniqueR cs = Unique cs × All UniqueNode cs @@ -255,7 +259,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where → (a : UniqueNode x) → (b : UniqueNode x) → a ≡ b - UniqueNode-deterministic {pnode _ cs} (unq a) (unq b) = Eq.cong unq (UniqueR-deterministic cs a b) + UniqueNode-deterministic {_ -< cs >- } (unq a) (unq b) = Eq.cong unq (UniqueR-deterministic cs a b) UniqueR-deterministic : ∀ (xs : List (FST A)) → (ua : UniqueR xs) @@ -278,20 +282,20 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where → UniqueR rs → UniqueR e ↪-preserves-unique impose-nothing ur-ls ur-rs = ur-rs - ↪-preserves-unique {pnode a as ∷ ls} {rs} (impose-step {e' = e'} ↝e' ↪e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = + ↪-preserves-unique {a -< as >- ∷ ls} {rs} (impose-step {e' = e'} ↝e' ↪e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = let ur-e' = ↝-preserves-unique a as rs e' ↝e' ur-as ur-rs ur-e = ↪-preserves-unique ↪e (u-ls , ur-ls) ur-e' in ur-e ↝-preserves-unique : ∀ (a : A) (ls rs : List (FST A)) (e : List (FST A)) - → pnode a ls ⊙ rs ↝ e + → a -< ls >- ⊙ rs ↝ e → UniqueR ls → UniqueR rs → UniqueR e ↝-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] - ↝-preserves-unique a ls (pnode .a bs ∷ rs) e@(pnode .a cs ∷ rs) (merge ↪e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) + ↝-preserves-unique a ls (.a -< bs >- ∷ rs) e@(.a -< cs >- ∷ rs) (merge ↪e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) = unique-ignores-children u-rs , unq (↪-preserves-unique ↪e ur-ls ur-bs) ∷ un-rs - ↝-preserves-unique a ls (pnode b bs ∷ rs) (pnode .b .bs ∷ cs) (skip a≠b ↝cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) + ↝-preserves-unique a ls (b -< bs >- ∷ rs) (.b -< .bs >- ∷ cs) (skip a≠b ↝cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) = ind a≠b (u-r ∷ u-rs) ↝cs ∷ u-cs , ur-bs ∷ un-cs where ur-cs = ↝-preserves-unique a ls rs cs ↝cs u-ls (u-rs , ur-rs) @@ -300,11 +304,11 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where ind : ∀ {a ls b bs cs rs} → ¬ (a ≡ b) - → Unique (pnode b bs ∷ rs) - → pnode a ls ⊙ rs ↝ cs - → All (_≉_ (pnode b bs)) cs - ind {a} {ls} {b} {bs} a≠b _ base = ≉-sym (pnode a ls) (pnode b bs) a≠b ∷ [] - ind {a} {_} {b} {bs} {pnode .a cs ∷ _} a≠b u-rs (merge _) = ≉-sym (pnode a cs) (pnode b bs) a≠b ∷ head (drop-second-Unique u-rs) + → Unique (b -< bs >- ∷ rs) + → a -< ls >- ⊙ rs ↝ cs + → All (_≉_ (b -< bs >-)) cs + ind {a} {ls} {b} {bs} a≠b _ base = ≉-sym (a -< ls >-) (b -< bs >-) a≠b ∷ [] + ind {a} {_} {b} {bs} {.a -< cs >- ∷ _} a≠b u-rs (merge _) = ≉-sym (a -< cs >-) (b -< bs >-) a≠b ∷ head (drop-second-Unique u-rs) ind a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ↝cs) = b≠b' ∷ ind a≠b (u-r ∷ u-rs) ↝cs -- Feature Structure Forest @@ -449,15 +453,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- Semantics ⟦_⟧ : SPL → Conf → Rose ∞ A - ⟦ r ◀ features ⟧ c = rose (r -< forget-uniqueness (⊕-all (map impl (select c features))) >-) - - -- We could avoid wrap and unwrap by defining our own intermediate tree structure - -- that does not reuse Artifact constructor. - -- unwrap : Rose A → Artifact A (Rose A) - -- unwrap (artifact a) = a - - -- wrap : Artifact A (Rose A) → Rose A - -- wrap a = artifact a + ⟦ r ◀ features ⟧ c = r -< forget-uniqueness (⊕-all (map impl (select c features))) >- open import Data.String using (String; _<+>_) open import Show.Lines From f677e5e9e7d12f3d7415e9191d3bdee0b90ad659 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 20 Mar 2024 15:35:24 +0100 Subject: [PATCH 32/43] simplify FST 'select' for hopefully simpler proofs --- src/Lang/FST.agda | 13 +++++++------ 1 file changed, 7 insertions(+), 6 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 6048d800..b5010f1a 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -371,11 +371,12 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where features : List Feature open SPL public - select : Conf → List Feature → List Feature - select c = filterᵇ (c ∘ name) - - -- forget-names : ∀ {N} → SPL N → List FSF - -- forget-names = map impl + select : Conf → List Feature → List FSF + select _ [] = [] + select c (f ∷ fs) = + if c (name f) + then impl f ∷ select c fs + else select c fs names : SPL → List F names = map name ∘ features @@ -453,7 +454,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- Semantics ⟦_⟧ : SPL → Conf → Rose ∞ A - ⟦ r ◀ features ⟧ c = r -< forget-uniqueness (⊕-all (map impl (select c features))) >- + ⟦ r ◀ features ⟧ c = r -< forget-uniqueness (⊕-all (select c features)) >- open import Data.String using (String; _<+>_) open import Show.Lines From ceb01aaf4e22b36a405710c4ba5503c57e47d9e8 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Wed, 20 Mar 2024 15:42:20 +0100 Subject: [PATCH 33/43] better names for Unique predicates --- src/Lang/FST.agda | 48 +++++++++++++++++++++++------------------------ 1 file changed, 24 insertions(+), 24 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index b5010f1a..9b8a8d1a 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -248,39 +248,39 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where drop-second-Unique ((_ ∷ pxs) ∷ _ ∷ zs) = pxs ∷ zs mutual - data UniqueNode : FST A → Set where - unq : ∀ {a cs} → UniqueR cs → UniqueNode (a -< cs >-) + data DescendantsUnique : FST A → Set where + unq : ∀ {a cs} → WellFormed cs → DescendantsUnique (a -< cs >-) - UniqueR : List (FST A) → Set - UniqueR cs = Unique cs × All UniqueNode cs + WellFormed : List (FST A) → Set + WellFormed cs = Unique cs × All DescendantsUnique cs mutual - UniqueNode-deterministic : ∀ {x : FST A} - → (a : UniqueNode x) - → (b : UniqueNode x) + DescendantsUnique-deterministic : ∀ {x : FST A} + → (a : DescendantsUnique x) + → (b : DescendantsUnique x) → a ≡ b - UniqueNode-deterministic {_ -< cs >- } (unq a) (unq b) = Eq.cong unq (UniqueR-deterministic cs a b) + DescendantsUnique-deterministic {_ -< cs >- } (unq a) (unq b) = Eq.cong unq (WellFormed-deterministic cs a b) - UniqueR-deterministic : ∀ (xs : List (FST A)) - → (ua : UniqueR xs) - → (ub : UniqueR xs) + WellFormed-deterministic : ∀ (xs : List (FST A)) + → (ua : WellFormed xs) + → (ub : WellFormed xs) → ua ≡ ub - UniqueR-deterministic [] ([] , []) ([] , []) = refl - UniqueR-deterministic (x ∷ xs) (a-x∉xs ∷ a-u-xs , a-ur-x ∷ a-ur-xs) (b-x∉xs ∷ b-u-xs , b-ur-x ∷ b-ur-xs) - with UniqueR-deterministic xs (a-u-xs , a-ur-xs) (b-u-xs , b-ur-xs) + WellFormed-deterministic [] ([] , []) ([] , []) = refl + WellFormed-deterministic (x ∷ xs) (a-x∉xs ∷ a-u-xs , a-ur-x ∷ a-ur-xs) (b-x∉xs ∷ b-u-xs , b-ur-x ∷ b-ur-xs) + with WellFormed-deterministic xs (a-u-xs , a-ur-xs) (b-u-xs , b-ur-xs) ... | eq rewrite (Eq.cong proj₁ eq) rewrite (Eq.cong proj₂ eq) - rewrite UniqueNode-deterministic a-ur-x b-ur-x + rewrite DescendantsUnique-deterministic a-ur-x b-ur-x rewrite ∉-deterministic xs a-x∉xs b-x∉xs = refl mutual ↪-preserves-unique : ∀ {ls rs e : List (FST A)} → ls + rs ↪ e - → UniqueR ls - → UniqueR rs - → UniqueR e + → WellFormed ls + → WellFormed rs + → WellFormed e ↪-preserves-unique impose-nothing ur-ls ur-rs = ur-rs ↪-preserves-unique {a -< as >- ∷ ls} {rs} (impose-step {e' = e'} ↝e' ↪e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = let ur-e' = ↝-preserves-unique a as rs e' ↝e' ur-as ur-rs @@ -289,9 +289,9 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where ↝-preserves-unique : ∀ (a : A) (ls rs : List (FST A)) (e : List (FST A)) → a -< ls >- ⊙ rs ↝ e - → UniqueR ls - → UniqueR rs - → UniqueR e + → WellFormed ls + → WellFormed rs + → WellFormed e ↝-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] ↝-preserves-unique a ls (.a -< bs >- ∷ rs) e@(.a -< cs >- ∷ rs) (merge ↪e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) = unique-ignores-children u-rs , unq (↪-preserves-unique ↪e ur-ls ur-bs) ∷ un-rs @@ -316,7 +316,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where constructor _⊚_ field trees : List (FST A) - valid : UniqueR trees + valid : WellFormed trees open FSF public ↝-append-strange : ∀ (l : FST A) (rs : List (FST A)) @@ -421,10 +421,10 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- Maybe we shouldnt prove for _≡_ but rather for a new eq relation -- that is weaker and ignores the typing. help : ∀ (ls : List (FST A)) - → (ur-ls : UniqueR ls) + → (ur-ls : WellFormed ls) → (deriv : ls + [] ↪ ls) → ↪-preserves-unique deriv ur-ls ([] , []) ≡ ur-ls - help ls ur-ls deriv = UniqueR-deterministic ls (↪-preserves-unique deriv ur-ls ([] , [])) ur-ls + help ls ur-ls deriv = WellFormed-deterministic ls (↪-preserves-unique deriv ur-ls ([] , [])) ur-ls assoc : Associative _≡_ _⊕_ assoc x y z = {!!} From 39938761b274095d0ef57c9d330cdcecd4f2081e Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 29 Mar 2024 14:20:54 +0100 Subject: [PATCH 34/43] FST variability lang --- src/Lang/FST.agda | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 9b8a8d1a..7b4412a2 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -23,8 +23,9 @@ open import Relation.Binary using (Decidable; DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) open import Framework.Annotation.Name using (Name) -open import Framework.Variants using (Rose; rose) +open import Framework.Variants using (Rose; rose; rose-leaf) open import Framework.Composition.FeatureAlgebra +open import Framework.VariabilityLanguage open import Construct.Artifact as At using () Conf : Set @@ -41,6 +42,7 @@ FST : 𝔼 FST = Rose ∞ pattern _-<_>- a cs = rose (a At.-< cs >-) +fst-leaf = rose-leaf induction : ∀ {A : 𝔸} {B : Set} → (A → List B → B) → FST A → B induction {A} {B} f n = go n [] where @@ -473,3 +475,21 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where show-Feature feature = do > show-F (name feature) <+> "∷" indent 2 (show-FSF (forget-uniqueness (impl feature))) + +-- Our framework does not allow constraints on the atom type A. +-- This demonstrates that FSTs are not totally generic. +-- DecidableEquality is a reasonable constraint though - or rather it is an axiom. +-- Maybe we could avoid this constraint by not comparing atoms A +-- but equipping each node with an additional ID and compare that instead? +-- That would allow merging nodes with unequal atoms though. +-- Maybe assuming decidable equality as an axiom here is just fine? +-- Not having decidable equality is far from practical. +module Framework (mkDec : (A : 𝔸) → DecidableEquality A) where + FSTL-𝔼 : 𝔼 + FSTL-𝔼 A = Impose.SPL (mkDec A) + + FSTL-Sem : 𝔼-Semantics (Rose ∞) Conf FSTL-𝔼 + FSTL-Sem {A} = Impose.⟦_⟧ (mkDec A) + + FSTL : VariabilityLanguage (Rose ∞) + FSTL = ⟪ FSTL-𝔼 , Conf , FSTL-Sem ⟫ From d681f6323df39421e71f6a95cd1cbcf32cd19442 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 29 Mar 2024 14:22:07 +0100 Subject: [PATCH 35/43] =?UTF-8?q?=C2=AC=CE=A3-syntax?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/Util/Existence.agda | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/src/Util/Existence.agda b/src/Util/Existence.agda index cbaa8791..b6513078 100644 --- a/src/Util/Existence.agda +++ b/src/Util/Existence.agda @@ -3,7 +3,8 @@ module Util.Existence where open import Agda.Primitive using (Level; _⊔_) -open import Data.Product using (Σ) +open import Data.Product using (Σ; ∄) +open import Relation.Nullary.Negation using (¬_) open import Size using (Size; SizeUniv) -- Existence of sizes @@ -27,3 +28,8 @@ open Eq using (_≡_; refl) ,-injectiveʳ : ∀ {A : Set} {i j : Size} {a b : A} → (i , a) ≡ (j , b) → a ≡ b ,-injectiveʳ refl = refl + +infix 2 ¬Σ-syntax +¬Σ-syntax : ∀ {a b} (A : Set a) → (A → Set b) → Set (a ⊔ b) +¬Σ-syntax _ p = ∄ p +syntax ¬Σ-syntax A (λ x → B) = ∄[ x ∈ A ] B From a685179d7a4cb62a2738e51dd91968c231425c39 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 29 Mar 2024 14:22:15 +0100 Subject: [PATCH 36/43] refactor: eta-reduce opt --- src/Lang/OC.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Lang/OC.lagda.md b/src/Lang/OC.lagda.md index 5de31049..c0113ff7 100644 --- a/src/Lang/OC.lagda.md +++ b/src/Lang/OC.lagda.md @@ -212,7 +212,7 @@ oc-leaf a = a -< [] >- -- alternative name that does not require writing tortoise shell braces opt : ∀ {i : Size} {A : 𝔸} → Option → OC i A → OC (↑ i) A -opt O = _❲_❳ O +opt = _❲_❳ singleton : ∀ {i : Size} {A : 𝔸} → A → OC i A → OC (↑ i) A singleton a e = a -< e ∷ [] >- From 8b4644313649494d6637cdd9376131264eb3a16e Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 29 Mar 2024 14:22:30 +0100 Subject: [PATCH 37/43] fixme: fst-to-oc --- src/Translation/Lang/FST-to-OC.agda | 186 ++++++++++++++++++++++++++++ 1 file changed, 186 insertions(+) create mode 100644 src/Translation/Lang/FST-to-OC.agda diff --git a/src/Translation/Lang/FST-to-OC.agda b/src/Translation/Lang/FST-to-OC.agda new file mode 100644 index 00000000..c3cfe406 --- /dev/null +++ b/src/Translation/Lang/FST-to-OC.agda @@ -0,0 +1,186 @@ +{-# OPTIONS --sized-types #-} + +open import Framework.Definitions +module Translation.Lang.FST-to-OC (F : 𝔽) where + +open import Size using (Size; ↑_; _⊔ˢ_; ∞) + +open import Data.List using (List; []; _∷_; map) +open import Relation.Nullary.Decidable using (yes; no; _because_; False) +open import Relation.Binary using (Decidable; DecidableEquality; Rel) +open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) + +open import Framework.Variants using (Rose; rose; Artifact∈ₛRose; rose-leaf) + +V = Rose ∞ +mkArtifact = Artifact∈ₛRose +Option = F + +open import Framework.VariabilityLanguage +open import Construct.Artifact as At using () +import Lang.FST as FST +-- open FST F using (FST; induction; fst-leaf) + +open import Framework.Relation.Expressiveness V + +-- module _ {A : 𝔸} (_==_ : DecidableEquality A) where +-- open FST.Impose F _==_ using (SPL; _◀_; Feature; _::_; FSF; _⊚_) + +-- -- translate-o : ∀ {i} → A → OC i A → Feature +-- -- translate-o a (b OC.-< cs >-) = {!!} +-- -- translate-o a (O OC.❲ e ❳) = {!!} + +-- -- left is base +-- mutual +-- embed : FST A → OC ∞ A +-- embed = induction OC._-<_>- + +-- merge-trees : F → List (OC ∞ A) → F → FST A → List (OC ∞ A) +-- merge-trees _ [] P t = P ❲ embed t ❳ ∷ [] +-- merge-trees O (a OC.-< as >- ∷ inter) P (b FST.-< bs >-) with a == b +-- ... | yes refl = a OC.-< merge O as P bs >- ∷ inter +-- ... | no nope = O ❲ a OC.-< as >- ❳ ∷ merge-trees O inter P (b FST.-< bs >-) +-- merge-trees O (Q ❲ e ❳ ∷ inter) P t = Q ❲ e ❳ ∷ merge-trees O inter P t + +-- {-# TERMINATING #-} +-- merge : F → List (OC ∞ A) → F → List (FST A) → List (OC ∞ A) +-- merge _ ls _ [] = ls +-- merge O ls P (t ∷ ts) = merge O (merge-trees O ls P t) P ts + +-- record Intermediate : Set where +-- constructor _:o:_ +-- field +-- name : F +-- children : List (OC ∞ A) + +-- translate' : List Intermediate → List (OC ∞ A) +-- translate' [] = [] +-- translate' ((_ :o: cs) ∷ []) = cs +-- translate' ((O :o: os) ∷ (P :o: ps) ∷ xs) = translate' ({!merge!} ∷ xs) + +-- fromFeature : Feature → Intermediate +-- fromFeature (O :: (ts ⊚ _)) = O :o: map embed ts + +-- start : List Feature → List Intermediate +-- start = map fromFeature + +-- translate : SPL → WFOC ∞ A +-- translate (a ◀ fs) = Root a (translate' (start fs)) + +module _ (mkDec : (A : 𝔸) → DecidableEquality A) where + data FeatureName : Set where + X : FeatureName + Y : FeatureName + + open import Data.Bool using (true; false; if_then_else_) + open import Data.Nat using (zero; suc; ℕ) + open import Data.Fin using (Fin; zero; suc) + open import Data.List.Relation.Unary.All using (All; []; _∷_) + open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) + open import Framework.VariantMap V ℕ + + import Lang.OC + open Lang.OC FeatureName using (OC; _❲_❳; WFOC; Root; opt; Configuration) + open Lang.OC.Sem FeatureName V mkArtifact + + open FST FeatureName using (_._) + open FST.Impose FeatureName (mkDec ℕ) using (SPL; _◀_; _::_; _⊚_; unq) renaming (⟦_⟧ to FST⟦_⟧) + open FST.Framework FeatureName mkDec + + open import Data.EqIndexedSet + open import Data.Empty using (⊥-elim) + + open import Data.Product using (_,_; ∃-syntax) + open import Util.Existence using (¬Σ-syntax) + + counterexample : VMap 3 + counterexample zero = rose-leaf 0 + counterexample (suc zero) = rose (0 At.-< rose (1 At.-< rose-leaf 2 ∷ [] >-) ∷ [] >-) + counterexample (suc (suc zero)) = rose (0 At.-< rose (1 At.-< rose-leaf 3 ∷ [] >-) ∷ [] >-) + counterexample (suc (suc (suc zero))) = rose (0 At.-< rose (1 At.-< rose-leaf 2 ∷ rose-leaf 3 ∷ [] >-) ∷ [] >-) + + -- illegal' : ∀ {i : Size} → ∄[ cs ∈ List (OC i ℕ) ] (⟦ Root zero cs ⟧ ≅ counterexample) + -- illegal' ([] , fst , snd) with snd (suc zero) + -- ... | () + -- illegal' (x ∷ fst , snd) = {!!} + + -- illegal' : ∀ {i : Size} → ∄[ e ∈ OC i ℕ ] (∃[ O ] (∃[ xs ] (⟦ Root zero (opt O e ∷ xs) ⟧ ≅ counterexample))) + -- illegal' x = {!!} + open import Relation.Nullary.Negation using (¬_) + + nodup : ∀ {i : Size} {A : 𝔸} (a : A) (x y : OC i A) (zs : List (OC i A)) + → ¬ (∀ (c : Configuration) → (⟦ Root a (x ∷ y ∷ zs) ⟧ c ≡ rose-leaf a)) + nodup a x y zs sure with sure (λ _ → true) + nodup a (O ❲ e ❳) y zs sure | asd = {!!} + -- nodup a (b OC.-< bs >-) y zs sure with sure (λ _ → true) + -- ... | () + -- nodup a (O ❲ e ❳) (b OC.-< bs >-) zs sure with sure (λ _ → false) + -- ... | () + -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] sure = {!!} + -- nodup a (O ❲ e ❳) (P ❲ f ❳) (z ∷ zs) sure = nodup a (P ❲ f ❳) z zs {!sure!} + -- nodup a (O ❲ e ❳) y zs (c , sure) with c O + -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] (c , sure) | false with c P + -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] (c , refl) | false | false = {!!} + -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] (c , sure) | false | true = {!!} + -- nodup a (O ❲ e ❳) (P ❲ f ❳) (z ∷ zs) (c , sure) | false = nodup a (P ❲ f ❳) z zs (c , sure) + -- nodup a (O ❲ e ❳) y zs (c , sure) | true = {!!} + + illegal : ∀ {i : Size} → ∄[ e ∈ WFOC i ℕ ] (⟦ e ⟧ ≅ counterexample) + -- root must be zero because it is always zero in counterexample + illegal (Root (suc i) cs , _ , ⊇) with ⊇ zero + ... | () + -- there must be a child because there are variants in counterexample with children below the root (e.g., suc zero) + illegal (Root zero [] , _ , ⊇) with ⊇ (suc zero) -- illegal' (cs , eq) + ... | () + -- there must be an option at the front because there are variants with zero children (e.g., zero) + illegal (Root zero (a OC.-< cs >- ∷ _) , _ , ⊇) with ⊇ zero + ... | () + -- there can not be any other children + illegal (Root zero (O ❲ e ❳ ∷ a OC.-< as >- ∷ xs) , ⊆ , ⊇) = {!!} + illegal (Root zero (O ❲ e ❳ ∷ P OC.❲ e' ❳ ∷ xs) , ⊆ , ⊇) = {!!} + -- e can be a chain of options but somewhen, there must be an artifact + illegal (Root zero (O ❲ e ❳ ∷ []) , ⊆ , ⊇) = {!!} + --illegal' (e , (O , xs , (⊆ , ⊇))) + + cef : SPL + cef = 0 ◀ ( + (X :: (1 . 2 . []) ⊚ ([] ∷ [] , unq ([] ∷ [] , unq ([] , []) ∷ []) ∷ [])) ∷ + (Y :: (1 . 3 . []) ⊚ ([] ∷ [] , unq ([] ∷ [] , unq ([] , []) ∷ []) ∷ [])) ∷ + []) + + cef-describes-counterexample : FST⟦ cef ⟧ ≅ counterexample + cef-describes-counterexample = ⊆[]→⊆ cef-⊆[] , ⊆[]→⊆ {f = fnoc} cef-⊇[] + where + conf : FST.Conf FeatureName → Fin 4 + conf c with c X | c Y + ... | false | false = zero + ... | false | true = suc (suc zero) + ... | true | false = suc zero + ... | true | true = suc (suc (suc zero)) + + cef-⊆[] : FST⟦ cef ⟧ ⊆[ conf ] counterexample + cef-⊆[] c with c X | c Y + ... | false | false = refl + ... | false | true = refl + ... | true | false = refl + ... | true | true = {!!} + + fnoc : Fin 4 → FST.Conf FeatureName + fnoc zero X = false + fnoc zero Y = false + fnoc (suc zero) X = true + fnoc (suc zero) Y = false + fnoc (suc (suc zero)) X = false + fnoc (suc (suc zero)) Y = true + fnoc (suc (suc (suc zero))) X = true + fnoc (suc (suc (suc zero))) Y = true + + cef-⊇[] : counterexample ⊆[ fnoc ] FST⟦ cef ⟧ + cef-⊇[] zero = refl + cef-⊇[] (suc zero) = refl + cef-⊇[] (suc (suc zero)) = refl + cef-⊇[] (suc (suc (suc zero))) = {!!} + + ouch : WFOCL ⋡ FSTL + ouch sure with sure cef + ... | e , e≣cef = {!!} From de74a516108486169671a4d2f97a77d6916c21ca Mon Sep 17 00:00:00 2001 From: pmbittner Date: Mon, 1 Apr 2024 20:11:44 +0200 Subject: [PATCH 38/43] simplified fsts to paper formulas --- src/Lang/FST.agda | 483 ++++++++++++++++++++++------------------------ 1 file changed, 233 insertions(+), 250 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 7b4412a2..53a723d6 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -6,7 +6,7 @@ module Lang.FST (F : 𝔽) where open import Data.Bool using (Bool; true; false; if_then_else_) open import Data.List using (List; []; _∷_; _∷ʳ_; _++_; foldr; map; filterᵇ; concat; reverse) -open import Data.List.Properties using (++-identityʳ) +open import Data.List.Properties using (++-identityˡ; ++-identityʳ) open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.List.Relation.Unary.All using (All; []; _∷_) renaming (map to map-all) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_; head) @@ -15,12 +15,15 @@ open import Data.Empty using (⊥; ⊥-elim) open import Data.Unit using (tt) open import Function using (_∘_) open import Level using (0ℓ) -open import Size using (∞) +open import Size using (Size; ↑_; ∞) + +open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) open import Relation.Nullary.Negation using (¬_) open import Relation.Nullary.Decidable using (yes; no; _because_; False) open import Relation.Binary using (Decidable; DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) +open Eq.≡-Reasoning open import Framework.Annotation.Name using (Name) open import Framework.Variants using (Rose; rose; rose-leaf) @@ -38,58 +41,58 @@ module TODO-MOVE-TO-AUX-OR-USE-STL where lem y ys (x ∷ xs) = Eq.cong (x ∷_) (lem y ys xs) open TODO-MOVE-TO-AUX-OR-USE-STL -FST : 𝔼 -FST = Rose ∞ +FST : Size → 𝔼 +FST i = Rose i pattern _-<_>- a cs = rose (a At.-< cs >-) fst-leaf = rose-leaf -induction : ∀ {A : 𝔸} {B : Set} → (A → List B → B) → FST A → B +induction : ∀ {A : 𝔸} {B : Set} → (A → List B → B) → FST ∞ A → B induction {A} {B} f n = go n [] where - go : FST A → List B → B + go : FST ∞ A → List B → B go (a -< [] >-) bs = f a (reverse bs) go (a -< c ∷ cs >-) bs = go (a -< cs >-) (go c [] ∷ bs) infix 15 _≈_ -_≈_ : ∀ {A} → Rel (FST A) 0ℓ +_≈_ : ∀ {A i} → Rel (FST i A) 0ℓ (a -< _ >-) ≈ (b -< _ >-) = a ≡ b -≈-sym : ∀ {A} → (a b : FST A) → a ≈ b → b ≈ a +≈-sym : ∀ {A i} → (a b : FST i A) → a ≈ b → b ≈ a ≈-sym (a -< _ >-) (.a -< _ >-) refl = refl infix 15 _≉_ -_≉_ : ∀ {A} → Rel (FST A) 0ℓ +_≉_ : ∀ {A i} → Rel (FST i A) 0ℓ a ≉ b = ¬ (a ≈ b) -≉-sym : ∀ {A} → (a b : FST A) → a ≉ b → b ≉ a +≉-sym : ∀ {A i} → (a b : FST i A) → a ≉ b → b ≉ a ≉-sym a b a≉b b≈a = a≉b (≈-sym b a b≈a) infix 15 _∈_ -_∈_ : ∀ {A} → FST A → List (FST A) → Set +_∈_ : ∀ {i A} → FST i A → List (FST i A) → Set x ∈ xs = Any (x ≈_) xs infix 15 _∉_ -_∉_ : ∀ {A} → FST A → List (FST A) → Set +_∉_ : ∀ {i A} → FST i A → List (FST i A) → Set x ∉ xs = All (x ≉_) xs -_⊑_ : ∀ {A} → (xs ys : List (FST A)) → Set --\squb= +_⊑_ : ∀ {i A} → (xs ys : List (FST i A)) → Set --\squb= xs ⊑ ys = All (_∈ ys) xs -_⋢_ : ∀ {A} → (xs ys : List (FST A)) → Set --\squb=n +_⋢_ : ∀ {i A} → (xs ys : List (FST i A)) → Set --\squb=n xs ⋢ ys = Any (_∉ ys) xs -Disjoint : ∀ {A} → (xs ys : List (FST A)) → Set --\squb=n +Disjoint : ∀ {i A} → (xs ys : List (FST i A)) → Set --\squb=n Disjoint xs ys = All (_∉ ys) xs -- identity of proofs open import Axioms.Extensionality using (extensionality) -≉-deterministic : ∀ {A} (x y : FST A) +≉-deterministic : ∀ {A} (x y : FST ∞ A) → (p₁ : x ≉ y) → (p₂ : x ≉ y) → p₁ ≡ p₂ ≉-deterministic (a -< _ >-) (b -< _ >-) p₁ p₂ = extensionality λ where refl → refl -∉-deterministic : ∀ {A} {x : FST A} (ys : List (FST A)) +∉-deterministic : ∀ {A} {x : FST ∞ A} (ys : List (FST ∞ A)) → (p₁ : x ∉ ys) → (p₂ : x ∉ ys) → p₁ ≡ p₂ @@ -99,25 +102,25 @@ open import Axioms.Extensionality using (extensionality) rewrite ∉-deterministic ys pa pb = refl -map-≉ : ∀ {A} {b xs} (ys : List (FST A)) (z : FST A) +map-≉ : ∀ {i} {A} {b xs} (ys : List (FST i A)) (z : FST (↑ i) A) → b -< xs >- ≉ z → b -< ys >- ≉ z map-≉ ys (z -< zs >-) z≉z refl = z≉z refl -map-∉ : ∀ {A} {b : A} {cs cs' xs : List (FST A)} +map-∉ : ∀ {i} {A} {b : A} {cs cs' : List (FST i A)} {xs : List (FST (↑ i) A)} → b -< cs >- ∉ xs → b -< cs' >- ∉ xs map-∉ [] = [] map-∉ {cs' = cs'} {xs = x ∷ xs} (px ∷ pxs) = map-≉ cs' x px ∷ map-∉ pxs -disjoint-[]ˡ : ∀ {A} (xs : List (FST A)) → Disjoint [] xs +disjoint-[]ˡ : ∀ {i A} (xs : List (FST i A)) → Disjoint [] xs disjoint-[]ˡ _ = [] -disjoint-[]ʳ : ∀ {A} (xs : List (FST A)) → Disjoint xs [] +disjoint-[]ʳ : ∀ {i A} (xs : List (FST i A)) → Disjoint xs [] disjoint-[]ʳ [] = [] disjoint-[]ʳ (x ∷ xs) = [] ∷ (disjoint-[]ʳ xs) -disjoint-grow : ∀ {A} (r : FST A) (rs ls : List (FST A)) +disjoint-grow : ∀ {i A} (r : FST i A) (rs ls : List (FST i A)) → Disjoint ls rs → r ∉ ls → Disjoint ls (r ∷ rs) @@ -125,14 +128,14 @@ disjoint-grow r rs [] _ _ = [] disjoint-grow r rs (l ∷ ls) (l∉rs ∷ d-ls-rs) (r≉l ∷ r∉ls) = (≉-sym r l r≉l ∷ l∉rs) ∷ disjoint-grow r rs ls d-ls-rs r∉ls -disjoint-shiftʳ : ∀ {A} (r : FST A) (rs ls : List (FST A)) +disjoint-shiftʳ : ∀ {i A} (r : FST i A) (rs ls : List (FST i A)) → Disjoint ls (r ∷ rs) → Disjoint ls (rs ++ r ∷ []) disjoint-shiftʳ r rs [] x = [] disjoint-shiftʳ r rs (l ∷ ls) ((l≉r ∷ l∉rs) ∷ d-ls-rrs) = step l r rs l≉r l∉rs ∷ disjoint-shiftʳ r rs ls d-ls-rrs where - step : ∀ {A} (x y : FST A) (zs : List (FST A)) + step : ∀ {i A} (x y : FST i A) (zs : List (FST i A)) → x ≉ y → x ∉ zs → x ∉ (zs ++ y ∷ []) @@ -141,211 +144,205 @@ disjoint-shiftʳ r rs (l ∷ ls) ((l≉r ∷ l∉rs) ∷ d-ls-rrs) -- the syntax used in the paper for paths infixr 5 _._ -_._ : ∀ {A} → A → (cs : List (FST A)) → List (FST A) +_._ : ∀ {A} → A → (cs : List (FST ∞ A)) → List (FST ∞ A) a . cs = a -< cs >- ∷ [] -- helper function when branching in paths -branches : ∀ {A} → List (List (FST A)) → List (FST A) +branches : ∀ {A} → List (List (FST ∞ A)) → List (FST ∞ A) branches = concat -mutual - infix 4 _⊙_↝_ - data _⊙_↝_ : ∀ {A} → FST A → List (FST A) → List (FST A) → Set where - base : ∀ {A} {l : FST A} - --------------- - → l ⊙ [] ↝ l ∷ [] - - merge : ∀ {A} {a : A} {as bs rs cs : List (FST A)} - → as + bs ↪ cs - ---------------------------------------------- - → a -< as >- ⊙ a -< bs >- ∷ rs ↝ a -< cs >- ∷ rs - - -- In the original work, skipped nodes were added to the left. - -- We add to the right here because it fits nicer with list construction _∷_ - -- Otherwise, we would have to backtrack when we found no match in rs. - skip : ∀ {A} {a r : FST A} {rs cs : List (FST A)} - → a ≉ r - → a ⊙ rs ↝ cs - ---------------------------------------------- - → a ⊙ r ∷ rs ↝ r ∷ cs - - -- This is basically just a fold on lists. Maybe we can simplify it accordingly. - infix 4 _+_↪_ - data _+_↪_ : ∀ {A} → List (FST A) → List (FST A) → List (FST A) → Set where - impose-nothing : ∀ {A} {rs : List (FST A)} - → [] + rs ↪ rs - - impose-step : ∀ {A} {l : FST A} {ls rs e e' : List (FST A)} - → l ⊙ rs ↝ e' - → ls + e' ↪ e - ---------------- - → l ∷ ls + rs ↪ e - -mutual - ↪-deterministic : ∀ {A} {fs rs e e' : List (FST A)} - → fs + rs ↪ e - → fs + rs ↪ e' - → e ≡ e' - ↪-deterministic impose-nothing impose-nothing = refl - ↪-deterministic (impose-step ↝x ↪x) (impose-step ↝y ↪y) rewrite ↝-deterministic ↝x ↝y | ↪-deterministic ↪x ↪y = refl - - ↝-deterministic : ∀ {A} {f : FST A} {rs e e' : List (FST A)} - → f ⊙ rs ↝ e - → f ⊙ rs ↝ e' - → e ≡ e' - ↝-deterministic base base = refl - ↝-deterministic (merge x) (merge y) rewrite ↪-deterministic x y = refl - ↝-deterministic (merge x) (skip a≠a y) = ⊥-elim (a≠a refl) - ↝-deterministic (skip a≠a x) (merge y) = ⊥-elim (a≠a refl) - ↝-deterministic (skip neq x) (skip neq' y) rewrite ↝-deterministic x y = refl - -↪-return : ∀ {A} {e ls rs : List (FST A)} - → ls + rs ↪ e - → ∃[ e ] (ls + rs ↪ e) -↪-return {e = e} ↪e = e , ↪e - -↝-return : ∀ {A} {l : FST A} {e rs : List (FST A)} - → l ⊙ rs ↝ e - → ∃[ e ] (l ⊙ rs ↝ e) -↝-return {e = e} ↝e = e , ↝e - -module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where - _==_ : Decidable (_≈_ {A}) +module _ {A : 𝔸} (_≟_ : DecidableEquality A) where + _==_ : ∀ {i} → Decidable (_≈_ {A} {i}) (a -< _ >-) == (b -< _ >-) = a ≟ b - childs : FST A → List (FST A) - childs (a -< as >-) = as + -- ≟-refl : ∀ (x : A) → x ≡ x + -- ≟-refl = {!!} mutual - ↪-total : ∀ (ls rs : List (FST A)) → ∃[ e ] (ls + rs ↪ e) - ↪-total [] rs = ↪-return impose-nothing - ↪-total (a -< as >- ∷ ls) rs = - let e' , ↝e' = ↝-total (a -< as >-) rs (↪-total as) - _ , ↪e = ↪-total ls e' - in ↪-return (impose-step ↝e' ↪e) - - ↝-total : ∀ (l : FST A) (rs : List (FST A)) - → ((rs' : List (FST A)) → ∃[ e ] (childs l + rs' ↪ e)) - → ∃[ e ] (l ⊙ rs ↝ e) - ↝-total l [] _ = ↝-return base - ↝-total (a -< as >-) (b -< bs >- ∷ rs) ↪-total-as with a ≟ b - ... | yes refl = - let cs , ↪cs = ↪-total-as bs - in ↝-return (merge ↪cs) - ... | no a≠b = - let cs , ↝cs = ↝-total (a -< as >-) rs ↪-total-as - in ↝-return (skip a≠b ↝cs) - - Unique : List (FST A) → Set - Unique = AllPairs _≉_ + infixr 5 _⊕_ + _⊕_ : ∀ {i} → List (FST i A) → List (FST i A) → List (FST i A) + l ⊕ [] = l + l ⊕ (h ∷ t) = (h ⊙ l) ⊕ t - unique-ignores-children : ∀ {a as bs rs} - → Unique (a -< as >- ∷ rs) - → Unique (a -< bs >- ∷ rs) - unique-ignores-children (x ∷ xs) = map-∉ x ∷ xs - drop-second-Unique : ∀ {x y zs} - → Unique (x ∷ y ∷ zs) - → Unique (x ∷ zs) - drop-second-Unique ((_ ∷ pxs) ∷ _ ∷ zs) = pxs ∷ zs + infixr 5 _⊙_ + _⊙_ : ∀ {i} → FST i A → List (FST i A) → List (FST i A) + l ⊙ [] = l ∷ [] + l ⊙ (h ∷ t) with l == h + ... | no _ = h ∷ (l ⊙ t) + a -< ca >- ⊙ (.a -< cb >- ∷ t) | yes refl = a -< ca ⊕ cb >- ∷ t + + Unique : ∀ {i} → List (FST i A) → Set + Unique = AllPairs _≉_ mutual - data DescendantsUnique : FST A → Set where - unq : ∀ {a cs} → WellFormed cs → DescendantsUnique (a -< cs >-) + WellFormed : ∀ {i} → FST i A → Set + WellFormed (_ -< cs >-) = AllWellFormed cs - WellFormed : List (FST A) → Set - WellFormed cs = Unique cs × All DescendantsUnique cs + AllWellFormed : ∀ {i} → List (FST i A) → Set + AllWellFormed cs = Unique cs × All WellFormed cs mutual - DescendantsUnique-deterministic : ∀ {x : FST A} - → (a : DescendantsUnique x) - → (b : DescendantsUnique x) + ⊕-wf : ∀ {i} {ls rs : List (FST i A)} + → AllWellFormed ls + → AllWellFormed rs + → AllWellFormed (ls ⊕ rs) + ⊕-wf ls-wf ([] , []) = ls-wf + ⊕-wf ls-wf (_ ∷ u-rs , du-r ∷ du-rs) = ⊕-wf (⊙-wf du-r ls-wf) (u-rs , du-rs) + + ⊙-wf : ∀ {i} {l : FST i A} {r : List (FST i A)} + → WellFormed l + → AllWellFormed r + → AllWellFormed (l ⊙ r) + ⊙-wf du-l ([] , []) = [] ∷ [] , du-l ∷ [] + ⊙-wf {_} {l} {h ∷ _} _ (_ ∷ _ , _ ∷ _) with l == h + ⊙-wf {_} {a -< ca >- } {(.a -< cb >-) ∷ t} wf-ca ( _ ∷ _ , wf-cb ∷ _) | yes refl with ⊕-wf wf-ca wf-cb + ⊙-wf _ (u-h ∷ u-t , _ ∷ du-t) | yes refl | wf-ca⊕cb + = (map-∉ u-h) ∷ u-t , wf-ca⊕cb ∷ du-t + ⊙-wf {_} {a -< ca >- } {b -< cb >- ∷ t} du-l (u-h ∷ u-t , du-h ∷ du-t) | no _ with ⊙-wf du-l (u-t , du-t) + ⊙-wf {_} {a -< ca >- } {b -< cb >- ∷ t} du-l (u-h ∷ u-t , du-h ∷ du-t) | no a≢b | u-rec , du-rec + = ind a≢b u-h ∷ u-rec , du-h ∷ du-rec + where + ind : ∀ {i} {b a} {cb ca : List (FST i A)} {t : List (FST (↑ i) A)} + → ¬ (a ≡ b) + → b -< cb >- ∉ t + → b -< cb >- ∉ ((a -< ca >-) ⊙ t) + ind {t = []} a≢b b∉t = (λ b≡a → a≢b (Eq.sym b≡a)) ∷ [] + ind {_} {_} {a} {_} {ca} {t ∷ ts} a≢b b∉t with (a -< ca >-) == t + ind {_} {_} {a} {cb} {ca} {(.a -< ct >-) ∷ ts} a≢b ( _ ∷ b∉ts) | yes refl = (λ b≡a → a≢b (Eq.sym b≡a)) ∷ b∉ts + ind {_} {_} {a} {cb} {ca} {( t -< ct >-) ∷ ts} a≢b (b≢t ∷ b∉ts) | no a≢t = b≢t ∷ (ind a≢b b∉ts) + + mutual + WellFormed-deterministic : ∀ {x : FST ∞ A} + → (a : WellFormed x) + → (b : WellFormed x) → a ≡ b - DescendantsUnique-deterministic {_ -< cs >- } (unq a) (unq b) = Eq.cong unq (WellFormed-deterministic cs a b) + WellFormed-deterministic {_ -< cs >- } a b = AllWellFormed-deterministic cs a b - WellFormed-deterministic : ∀ (xs : List (FST A)) - → (ua : WellFormed xs) - → (ub : WellFormed xs) + AllWellFormed-deterministic : ∀ (xs : List (FST ∞ A)) + → (ua : AllWellFormed xs) + → (ub : AllWellFormed xs) → ua ≡ ub - WellFormed-deterministic [] ([] , []) ([] , []) = refl - WellFormed-deterministic (x ∷ xs) (a-x∉xs ∷ a-u-xs , a-ur-x ∷ a-ur-xs) (b-x∉xs ∷ b-u-xs , b-ur-x ∷ b-ur-xs) - with WellFormed-deterministic xs (a-u-xs , a-ur-xs) (b-u-xs , b-ur-xs) + AllWellFormed-deterministic [] ([] , []) ([] , []) = refl + AllWellFormed-deterministic (x ∷ xs) (a-x∉xs ∷ a-u-xs , a-ur-x ∷ a-ur-xs) (b-x∉xs ∷ b-u-xs , b-ur-x ∷ b-ur-xs) + with AllWellFormed-deterministic xs (a-u-xs , a-ur-xs) (b-u-xs , b-ur-xs) ... | eq rewrite (Eq.cong proj₁ eq) rewrite (Eq.cong proj₂ eq) - rewrite DescendantsUnique-deterministic a-ur-x b-ur-x + rewrite WellFormed-deterministic a-ur-x b-ur-x rewrite ∉-deterministic xs a-x∉xs b-x∉xs = refl - mutual - ↪-preserves-unique : ∀ {ls rs e : List (FST A)} - → ls + rs ↪ e - → WellFormed ls - → WellFormed rs - → WellFormed e - ↪-preserves-unique impose-nothing ur-ls ur-rs = ur-rs - ↪-preserves-unique {a -< as >- ∷ ls} {rs} (impose-step {e' = e'} ↝e' ↪e) (u-l ∷ u-ls , unq ur-as ∷ ur-ls) ur-rs = - let ur-e' = ↝-preserves-unique a as rs e' ↝e' ur-as ur-rs - ur-e = ↪-preserves-unique ↪e (u-ls , ur-ls) ur-e' - in ur-e - - ↝-preserves-unique : ∀ (a : A) (ls rs : List (FST A)) (e : List (FST A)) - → a -< ls >- ⊙ rs ↝ e - → WellFormed ls - → WellFormed rs - → WellFormed e - ↝-preserves-unique _ _ _ _ base ur-ls _ = [] ∷ [] , (unq ur-ls) ∷ [] - ↝-preserves-unique a ls (.a -< bs >- ∷ rs) e@(.a -< cs >- ∷ rs) (merge ↪e) ur-ls (u-rs , (unq ur-bs) ∷ un-rs) - = unique-ignores-children u-rs , unq (↪-preserves-unique ↪e ur-ls ur-bs) ∷ un-rs - ↝-preserves-unique a ls (b -< bs >- ∷ rs) (.b -< .bs >- ∷ cs) (skip a≠b ↝cs) u-ls (u-r ∷ u-rs , ur-bs ∷ ur-rs) - = ind a≠b (u-r ∷ u-rs) ↝cs ∷ u-cs , ur-bs ∷ un-cs - where - ur-cs = ↝-preserves-unique a ls rs cs ↝cs u-ls (u-rs , ur-rs) - u-cs = proj₁ ur-cs - un-cs = proj₂ ur-cs + ⊙-stranger : ∀ {i} (l : FST i A) (rs : List (FST i A)) + → l ∉ rs + → l ⊙ rs ≡ rs ∷ʳ l + ⊙-stranger l [] _ = refl + ⊙-stranger l (r ∷ rs) (l≢r ∷ l∉rs) with l == r -- TODO: Is there an easier way to tell Agda that we already know l ≢ r? + ... | yes l≡r = ⊥-elim (l≢r l≡r) + ... | no _ = Eq.cong (r ∷_) (⊙-stranger l rs l∉rs) + + ⊕-strangers : ∀ {i} (ls rs : List (FST i A)) + → Unique rs + → Disjoint rs ls + → ls ⊕ rs ≡ ls ++ rs + ⊕-strangers ls [] _ _ rewrite ++-identityʳ ls = refl + ⊕-strangers ls (r ∷ rs) (r∉rs ∷ u-rs) (r∉ls ∷ d-ls-rs) +-- Goal: (r ⊙ ls) ⊕ rs ≡ ls ++ r ∷ rs + rewrite (Eq.sym (lem r rs ls)) +-- Goal: (r ⊙ ls) ⊕ rs ≡ (ls ++ r ∷ []) ++ rs + rewrite ⊙-stranger r ls r∉ls +-- Goal: (ls ++ r ∷ []) ⊕ rs ≡ (ls ++ r ∷ []) ++ rs + = ⊕-strangers (ls ++ r ∷ []) rs u-rs (disjoint-shiftʳ r ls rs (disjoint-grow r ls rs d-ls-rs r∉rs)) + + ⊕-idˡ : + ∀ {i} (rs : List (FST i A)) + → Unique rs + → [] ⊕ rs ≡ rs + ⊕-idˡ rs u-rs = ⊕-strangers [] rs u-rs (disjoint-[]ʳ rs) + + -- A proof that all FSTs xs are already imposed into another list of FSTs ys. + data _lies-in_ : ∀ {i} → List (FST i A) → List (FST i A) → Set where + lempty : ∀ {i} {xs : List (FST i A)} + ------------- + → [] lies-in xs + + lstep-here : ∀ {i} {a b : A} {as bs : List (FST i A)} {xs ys : List (FST (↑ i) A)} + → a ≡ b + → as lies-in bs + → xs lies-in ys + ---..................---------------------- + → (a -< as >- ∷ xs) lies-in (b -< bs >- ∷ ys) - ind : ∀ {a ls b bs cs rs} - → ¬ (a ≡ b) - → Unique (b -< bs >- ∷ rs) - → a -< ls >- ⊙ rs ↝ cs - → All (_≉_ (b -< bs >-)) cs - ind {a} {ls} {b} {bs} a≠b _ base = ≉-sym (a -< ls >-) (b -< bs >-) a≠b ∷ [] - ind {a} {_} {b} {bs} {.a -< cs >- ∷ _} a≠b u-rs (merge _) = ≉-sym (a -< cs >-) (b -< bs >-) a≠b ∷ head (drop-second-Unique u-rs) - ind a≠b ((b≠b' ∷ u-r) ∷ _ ∷ u-rs) (skip a≠b' ↝cs) = b≠b' ∷ ind a≠b (u-r ∷ u-rs) ↝cs + lstep-there : ∀ {i} {x y : FST i A} {xs ys : List (FST i A)} + → x ≉ y + → (x ∷ xs) lies-in ys + ------------------------- + → (x ∷ xs) lies-in (y ∷ ys) + + _slice-of_ : ∀ {i} → FST i A → FST i A → Set + x slice-of y = (x ∷ []) lies-in (y ∷ []) + + _slice-within_ : ∀ {i} → FST i A → List (FST i A) → Set + x slice-within ys = (x ∷ []) lies-in ys + + lies-in-refl : ∀ {i} → (xs : List (FST i A)) → xs lies-in xs + lies-in-refl [] = lempty + lies-in-refl ((a -< as >-) ∷ xs) = lstep-here refl (lies-in-refl as) (lies-in-refl xs) + + slice-prop : ∀ {i} {xs ys : List (FST i A)} (zs : List (FST i A)) + → xs lies-in ys + → xs lies-in (ys ⊕ zs) + slice-prop zs lempty = lempty + slice-prop {xs = a -< as >- ∷ xs} {ys = .a -< bs >- ∷ ys} zs (lstep-here refl as-lies-in-bs xs-lies-in-ys) = {!lstep-here!} + slice-prop zs (lstep-there x x₁) = {!!} + + slice-concat : ∀ {i} {x : FST i A} {ys : List (FST i A)} (xs : List (FST i A)) + → x slice-within ys + → (x ∷ xs) lies-in (ys ⊕ xs) + slice-concat = {!!} + + -- mutual + -- ⊕-makes-slicesˡ : ∀ {i} (xs ys : List (FST i A)) + -- → xs lies-in (ys ⊕ xs) + -- ⊕-makes-slicesˡ [] ys = lempty + -- ⊕-makes-slicesˡ (x ∷ xs) ys = slice-concat xs (⊙-makes-slice-head x ys) + + -- ⊕-makes-slicesʳ : ∀ {i} (xs ys : List (FST i A)) + -- → xs lies-in (xs ⊕ ys) + -- ⊕-makes-slicesʳ xs [] = lies-in-refl xs + -- ⊕-makes-slicesʳ xs (y ∷ ys) = slice-prop ys (⊙-makes-slice-tail y xs) + + -- ⊙-makes-slice-tail : ∀ {i} (x : FST i A) (ys : List (FST i A)) + -- → ys lies-in (x ⊙ ys) + -- ⊙-makes-slice-tail x [] = lempty + -- ⊙-makes-slice-tail (a -< cs >-) ((b -< bs >-) ∷ ys) with a ≟ b + -- ... | yes refl = lstep-here refl (⊕-makes-slicesˡ bs cs) (lies-in-refl ys) + -- ... | no _ = lstep-here refl (lies-in-refl bs) (⊙-makes-slice-tail (a -< cs >-) ys) + + -- ⊙-makes-slice-head : ∀ {i} (x : FST i A) (ys : List (FST i A)) + -- → x slice-within (x ⊙ ys) + -- ⊙-makes-slice-head (a -< cs >-) [] = lies-in-refl (a -< cs >- ∷ []) + -- ⊙-makes-slice-head (a -< cs >-) ((b -< bs >-) ∷ ys) with a ≟ b + -- ... | yes refl = lstep-here refl (⊕-makes-slicesʳ cs bs) lempty + -- ... | no a≠b = lstep-there a≠b (⊙-makes-slice-head (a -< cs >-) ys) + + ⊕-idem : ∀ {i} (xs ys : List (FST i A)) + → AllWellFormed xs + → AllWellFormed ys + → ys ⊕ xs ⊕ ys ≡ xs ⊕ ys + ⊕-idem xs [] (u-xs , _) _ = ⊕-idˡ xs u-xs + ⊕-idem [] (y ∷ ys) ([] , []) (y∉ys ∷ u-ys , wf-y ∷ wf-ys) = {!!} + ⊕-idem (x ∷ xs) (y ∷ ys) xs-wf ys-wf = {!!} -- Feature Structure Forest record FSF : Set where constructor _⊚_ field - trees : List (FST A) - valid : WellFormed trees + trees : List (FST ∞ A) + valid : AllWellFormed trees open FSF public - ↝-append-strange : ∀ (l : FST A) (rs : List (FST A)) - → l ∉ rs - → l ⊙ rs ↝ rs ∷ʳ l - ↝-append-strange l [] _ = base - ↝-append-strange l (r ∷ rs) (l≠r ∷ l∉rs) = skip l≠r (↝-append-strange l rs l∉rs) - - ↪-append-strangers : ∀ (ls rs : List (FST A)) - → Unique ls - → Disjoint ls rs - → ls + rs ↪ rs ++ ls - ↪-append-strangers [] rs _ _ rewrite ++-identityʳ rs = impose-nothing - ↪-append-strangers (l ∷ ls) rs (l∉ls ∷ u-ls) (l∉rs ∷ d-ls-rs) - rewrite (Eq.sym (lem l ls rs)) - with ↝-append-strange l rs l∉rs - ... | x - = impose-step x (↪-append-strangers ls (rs ++ l ∷ []) u-ls - (disjoint-shiftʳ l rs ls (disjoint-grow l rs ls d-ls-rs l∉ls))) - - impose-nothing-r : - ∀ (ls : List (FST A)) - → Unique ls - → ls + [] ↪ ls - impose-nothing-r ls u-ls = ↪-append-strangers ls [] u-ls (disjoint-[]ʳ ls) - - forget-uniqueness : FSF → List (FST A) + forget-uniqueness : FSF → List (FST ∞ A) forget-uniqueness = trees {- @@ -380,64 +377,50 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where then impl f ∷ select c fs else select c fs - names : SPL → List F - names = map name ∘ features - ---- Algebra ---- - open import Algebra.Definitions using (LeftIdentity; RightIdentity; Associative; Congruent₂) - open Eq.≡-Reasoning 𝟘 : FSF 𝟘 = [] ⊚ ([] , []) - infixr 7 _⊕_ - _⊕_ : FSF → FSF → FSF - (l ⊚ u-l) ⊕ (r ⊚ u-r) = - let e , ↪e = ↪-total l r - u-e = ↪-preserves-unique ↪e u-l u-r - in e ⊚ u-e - - ⊕-all : List FSF → FSF - ⊕-all = foldr _⊕_ 𝟘 - - l-id : LeftIdentity _≡_ 𝟘 _⊕_ - l-id _ = refl - - r-id : RightIdentity _≡_ 𝟘 _⊕_ - r-id (xs ⊚ (u-xs , ur-xs)) - -- Let's see what ⊕ does - with ↪-total xs [] - -- it computes some result 'e' and a derivation 'deriv' - ... | (e , deriv) - -- However, we know by impose-nothing-r that we can derive - -- 'xs' itself as result. - -- By determinism, we know that there can only be one derivation - -- so e = xs. - -- (We can't do a rewrite here for some reason so we stick to good old "with".) - with ↪-deterministic deriv (impose-nothing-r xs u-xs) - ... | refl = Eq.cong (xs ⊚_) (help xs (u-xs , ur-xs) deriv) - where - -- lastly, we have to prove that the typing is also unique but that is actually - -- irrelevant. Maybe we can avoid this proof somehow? - -- Its never needed and just an artifical problem. - -- Maybe we shouldnt prove for _≡_ but rather for a new eq relation - -- that is weaker and ignores the typing. - help : ∀ (ls : List (FST A)) - → (ur-ls : WellFormed ls) - → (deriv : ls + [] ↪ ls) - → ↪-preserves-unique deriv ur-ls ([] , []) ≡ ur-ls - help ls ur-ls deriv = WellFormed-deterministic ls (↪-preserves-unique deriv ur-ls ([] , [])) ur-ls - - assoc : Associative _≡_ _⊕_ - assoc x y z = {!!} - - cong : Congruent₂ _≡_ _⊕_ + infixr 7 _⊛_ + _⊛_ : FSF → FSF → FSF + (l ⊚ u-l) ⊛ (r ⊚ u-r) = (l ⊕ r) ⊚ (⊕-wf u-l u-r) + + ⊛-all : List FSF → FSF + ⊛-all = foldr _⊛_ 𝟘 + + cong-app₂ : ∀ {A C : Set} {T : A → Set} {x y : A} {tx : T x} {ty : T y} + → (f : (a : A) → T a → C) + → x ≡ y + → (∀ (a : A) (t u : T a) → t ≡ u) + → f x tx ≡ f y ty + cong-app₂ {y = y} {tx = tx} {ty = ty} f refl T-cong = Eq.cong (f y) (T-cong y tx ty) + + l-id : LeftIdentity _≡_ 𝟘 _⊛_ + l-id (ls ⊚ (u-ls , du-ls)) = cong-app₂ _⊚_ (⊕-idˡ ls u-ls) AllWellFormed-deterministic + + r-id : RightIdentity _≡_ 𝟘 _⊛_ + r-id (xs ⊚ (u-xs , ur-xs)) = refl + + -- ⊛ is not associative because + -- ⊕ is not associative because + -- the order in which children appear below their parents + -- is swapped. + -- Example: + -- X :: a -< b >- + -- Y :: a -< c >- + -- X ⊕ Y = a -< b , c >- + -- Y ⊕ X = a -< c , b >- + assoc : Associative _≡_ _⊛_ + assoc (x ⊚ x-wf) (y ⊚ y-wf) (z ⊚ z-wf) = {!!} + + cong : Congruent₂ _≡_ _⊛_ cong refl refl = refl - idem : ∀ (i₁ i₂ : FSF) → i₂ ⊕ i₁ ⊕ i₂ ≡ i₁ ⊕ i₂ + idem : ∀ (x y : FSF) → y ⊛ x ⊛ y ≡ x ⊛ y idem = {!!} - FST-is-FeatureAlgebra : FeatureAlgebra FSF _⊕_ 𝟘 + FST-is-FeatureAlgebra : FeatureAlgebra FSF _⊛_ 𝟘 FST-is-FeatureAlgebra = record { monoid = record { isSemigroup = record @@ -456,19 +439,19 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- Semantics ⟦_⟧ : SPL → Conf → Rose ∞ A - ⟦ r ◀ features ⟧ c = r -< forget-uniqueness (⊕-all (select c features)) >- + ⟦ r ◀ features ⟧ c = r -< forget-uniqueness (⊛-all (select c features)) >- open import Data.String using (String; _<+>_) open import Show.Lines module Show (show-F : F → String) (show-A : A → String) where mutual - show-FST : FST A → Lines + show-FST : FST ∞ A → Lines show-FST = induction λ a children → do > show-A a indent 2 (lines children) - show-FSF : List (FST A) → Lines + show-FSF : List (FST ∞ A) → Lines show-FSF roots = lines (map show-FST roots) show-Feature : Feature → Lines @@ -486,10 +469,10 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where -- Not having decidable equality is far from practical. module Framework (mkDec : (A : 𝔸) → DecidableEquality A) where FSTL-𝔼 : 𝔼 - FSTL-𝔼 A = Impose.SPL (mkDec A) + FSTL-𝔼 A = SPL (mkDec A) FSTL-Sem : 𝔼-Semantics (Rose ∞) Conf FSTL-𝔼 - FSTL-Sem {A} = Impose.⟦_⟧ (mkDec A) + FSTL-Sem {A} = ⟦_⟧ (mkDec A) FSTL : VariabilityLanguage (Rose ∞) FSTL = ⟪ FSTL-𝔼 , Conf , FSTL-Sem ⟫ From f835d24c4d9e848982079ae6bba2c4916a1f6afa Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 5 Apr 2024 21:55:41 +0200 Subject: [PATCH 39/43] documented todo for index equivalence --- src/Framework/Relation/Index.lagda.md | 1 + 1 file changed, 1 insertion(+) diff --git a/src/Framework/Relation/Index.lagda.md b/src/Framework/Relation/Index.lagda.md index 192f67dc..dc368d97 100644 --- a/src/Framework/Relation/Index.lagda.md +++ b/src/Framework/Relation/Index.lagda.md @@ -12,6 +12,7 @@ open import Framework.VariabilityLanguage ## Equivalence of Indices Two indices are equivalent for an expression when they produce the same output for all expressions. +We can actually generalize this to index equivalence on indexed sets (TODO). ```agda 𝕃 = VariabilityLanguage V From 503d3a48b2133bad530ae61812f44e5e65a1e3c4 Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 5 Apr 2024 21:56:02 +0200 Subject: [PATCH 40/43] progress on FSTs --- src/Lang/FST.agda | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 53a723d6..87add532 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -151,7 +151,7 @@ a . cs = a -< cs >- ∷ [] branches : ∀ {A} → List (List (FST ∞ A)) → List (FST ∞ A) branches = concat -module _ {A : 𝔸} (_≟_ : DecidableEquality A) where +module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where _==_ : ∀ {i} → Decidable (_≈_ {A} {i}) (a -< _ >-) == (b -< _ >-) = a ≟ b @@ -402,8 +402,8 @@ module _ {A : 𝔸} (_≟_ : DecidableEquality A) where r-id : RightIdentity _≡_ 𝟘 _⊛_ r-id (xs ⊚ (u-xs , ur-xs)) = refl - -- ⊛ is not associative because - -- ⊕ is not associative because + -- ⊛ is not commutative because + -- ⊕ is not commutative because -- the order in which children appear below their parents -- is swapped. -- Example: @@ -469,10 +469,10 @@ module _ {A : 𝔸} (_≟_ : DecidableEquality A) where -- Not having decidable equality is far from practical. module Framework (mkDec : (A : 𝔸) → DecidableEquality A) where FSTL-𝔼 : 𝔼 - FSTL-𝔼 A = SPL (mkDec A) + FSTL-𝔼 A = Impose.SPL (mkDec A) FSTL-Sem : 𝔼-Semantics (Rose ∞) Conf FSTL-𝔼 - FSTL-Sem {A} = ⟦_⟧ (mkDec A) + FSTL-Sem {A} = Impose.⟦_⟧ (mkDec A) FSTL : VariabilityLanguage (Rose ∞) FSTL = ⟪ FSTL-𝔼 , Conf , FSTL-Sem ⟫ From bc7de9daaa2b8f9adc613fdbede55d7b7cbdb04b Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 5 Apr 2024 22:06:09 +0200 Subject: [PATCH 41/43] integrate FSTs to new framework --- src/Lang/FST.agda | 107 +++++++++++++++++++++++----------------------- 1 file changed, 53 insertions(+), 54 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index 87add532..b9d9857b 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -2,6 +2,12 @@ {-# OPTIONS --sized-types #-} open import Framework.Definitions + +{- +This module formalizes feature structure trees. +We formalized the language, its semantics, and the typing to disallow duplicate neighbors. +We also prove that FSTs are a feature algebra but the proof is work in progress. +-} module Lang.FST (F : 𝔽) where open import Data.Bool using (Bool; true; false; if_then_else_) @@ -47,7 +53,7 @@ FST i = Rose i pattern _-<_>- a cs = rose (a At.-< cs >-) fst-leaf = rose-leaf -induction : ∀ {A : 𝔸} {B : Set} → (A → List B → B) → FST ∞ A → B +induction : ∀ {A : 𝔸} {B : Set} → (atoms A → List B → B) → FST ∞ A → B induction {A} {B} f n = go n [] where go : FST ∞ A → List B → B go (a -< [] >-) bs = f a (reverse bs) @@ -107,7 +113,7 @@ map-≉ : ∀ {i} {A} {b xs} (ys : List (FST i A)) (z : FST (↑ i) A) → b -< ys >- ≉ z map-≉ ys (z -< zs >-) z≉z refl = z≉z refl -map-∉ : ∀ {i} {A} {b : A} {cs cs' : List (FST i A)} {xs : List (FST (↑ i) A)} +map-∉ : ∀ {i} {A : 𝔸} {b : atoms A} {cs cs' : List (FST i A)} {xs : List (FST (↑ i) A)} → b -< cs >- ∉ xs → b -< cs' >- ∉ xs map-∉ [] = [] @@ -144,15 +150,20 @@ disjoint-shiftʳ r rs (l ∷ ls) ((l≉r ∷ l∉rs) ∷ d-ls-rrs) -- the syntax used in the paper for paths infixr 5 _._ -_._ : ∀ {A} → A → (cs : List (FST ∞ A)) → List (FST ∞ A) +_._ : ∀ {A : 𝔸} → atoms A → (cs : List (FST ∞ A)) → List (FST ∞ A) a . cs = a -< cs >- ∷ [] -- helper function when branching in paths branches : ∀ {A} → List (List (FST ∞ A)) → List (FST ∞ A) branches = concat -module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where - _==_ : ∀ {i} → Decidable (_≈_ {A} {i}) +module Impose (AtomSet : 𝔸) where + FSTA : Size → Set + FSTA i = FST i AtomSet + A = atoms AtomSet + _≟_ = proj₂ AtomSet + + _==_ : ∀ {i} → Decidable (_≈_ {AtomSet} {i}) (a -< _ >-) == (b -< _ >-) = a ≟ b -- ≟-refl : ∀ (x : A) → x ≡ x @@ -160,37 +171,37 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where mutual infixr 5 _⊕_ - _⊕_ : ∀ {i} → List (FST i A) → List (FST i A) → List (FST i A) + _⊕_ : ∀ {i} → List (FSTA i) → List (FSTA i) → List (FSTA i) l ⊕ [] = l l ⊕ (h ∷ t) = (h ⊙ l) ⊕ t infixr 5 _⊙_ - _⊙_ : ∀ {i} → FST i A → List (FST i A) → List (FST i A) + _⊙_ : ∀ {i} → FSTA i → List (FSTA i) → List (FSTA i) l ⊙ [] = l ∷ [] l ⊙ (h ∷ t) with l == h ... | no _ = h ∷ (l ⊙ t) a -< ca >- ⊙ (.a -< cb >- ∷ t) | yes refl = a -< ca ⊕ cb >- ∷ t - Unique : ∀ {i} → List (FST i A) → Set + Unique : ∀ {i} → List (FSTA i) → Set Unique = AllPairs _≉_ mutual - WellFormed : ∀ {i} → FST i A → Set + WellFormed : ∀ {i} → FSTA i → Set WellFormed (_ -< cs >-) = AllWellFormed cs - AllWellFormed : ∀ {i} → List (FST i A) → Set + AllWellFormed : ∀ {i} → List (FSTA i) → Set AllWellFormed cs = Unique cs × All WellFormed cs mutual - ⊕-wf : ∀ {i} {ls rs : List (FST i A)} + ⊕-wf : ∀ {i} {ls rs : List (FSTA i)} → AllWellFormed ls → AllWellFormed rs → AllWellFormed (ls ⊕ rs) ⊕-wf ls-wf ([] , []) = ls-wf ⊕-wf ls-wf (_ ∷ u-rs , du-r ∷ du-rs) = ⊕-wf (⊙-wf du-r ls-wf) (u-rs , du-rs) - ⊙-wf : ∀ {i} {l : FST i A} {r : List (FST i A)} + ⊙-wf : ∀ {i} {l : FSTA i} {r : List (FSTA i)} → WellFormed l → AllWellFormed r → AllWellFormed (l ⊙ r) @@ -203,7 +214,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where ⊙-wf {_} {a -< ca >- } {b -< cb >- ∷ t} du-l (u-h ∷ u-t , du-h ∷ du-t) | no a≢b | u-rec , du-rec = ind a≢b u-h ∷ u-rec , du-h ∷ du-rec where - ind : ∀ {i} {b a} {cb ca : List (FST i A)} {t : List (FST (↑ i) A)} + ind : ∀ {i} {b a} {cb ca : List (FSTA i)} {t : List (FSTA (↑ i))} → ¬ (a ≡ b) → b -< cb >- ∉ t → b -< cb >- ∉ ((a -< ca >-) ⊙ t) @@ -213,13 +224,13 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where ind {_} {_} {a} {cb} {ca} {( t -< ct >-) ∷ ts} a≢b (b≢t ∷ b∉ts) | no a≢t = b≢t ∷ (ind a≢b b∉ts) mutual - WellFormed-deterministic : ∀ {x : FST ∞ A} + WellFormed-deterministic : ∀ {x : FSTA ∞} → (a : WellFormed x) → (b : WellFormed x) → a ≡ b WellFormed-deterministic {_ -< cs >- } a b = AllWellFormed-deterministic cs a b - AllWellFormed-deterministic : ∀ (xs : List (FST ∞ A)) + AllWellFormed-deterministic : ∀ (xs : List (FSTA ∞)) → (ua : AllWellFormed xs) → (ub : AllWellFormed xs) → ua ≡ ub @@ -233,7 +244,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where rewrite ∉-deterministic xs a-x∉xs b-x∉xs = refl - ⊙-stranger : ∀ {i} (l : FST i A) (rs : List (FST i A)) + ⊙-stranger : ∀ {i} (l : FSTA i) (rs : List (FSTA i)) → l ∉ rs → l ⊙ rs ≡ rs ∷ʳ l ⊙-stranger l [] _ = refl @@ -241,7 +252,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where ... | yes l≡r = ⊥-elim (l≢r l≡r) ... | no _ = Eq.cong (r ∷_) (⊙-stranger l rs l∉rs) - ⊕-strangers : ∀ {i} (ls rs : List (FST i A)) + ⊕-strangers : ∀ {i} (ls rs : List (FSTA i)) → Unique rs → Disjoint rs ls → ls ⊕ rs ≡ ls ++ rs @@ -255,78 +266,78 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where = ⊕-strangers (ls ++ r ∷ []) rs u-rs (disjoint-shiftʳ r ls rs (disjoint-grow r ls rs d-ls-rs r∉rs)) ⊕-idˡ : - ∀ {i} (rs : List (FST i A)) + ∀ {i} (rs : List (FSTA i)) → Unique rs → [] ⊕ rs ≡ rs ⊕-idˡ rs u-rs = ⊕-strangers [] rs u-rs (disjoint-[]ʳ rs) -- A proof that all FSTs xs are already imposed into another list of FSTs ys. - data _lies-in_ : ∀ {i} → List (FST i A) → List (FST i A) → Set where - lempty : ∀ {i} {xs : List (FST i A)} + data _lies-in_ : ∀ {i} → List (FSTA i) → List (FSTA i) → Set where + lempty : ∀ {i} {xs : List (FSTA i)} ------------- → [] lies-in xs - lstep-here : ∀ {i} {a b : A} {as bs : List (FST i A)} {xs ys : List (FST (↑ i) A)} + lstep-here : ∀ {i} {a b : A} {as bs : List (FSTA i)} {xs ys : List (FSTA (↑ i))} → a ≡ b → as lies-in bs → xs lies-in ys ---..................---------------------- → (a -< as >- ∷ xs) lies-in (b -< bs >- ∷ ys) - lstep-there : ∀ {i} {x y : FST i A} {xs ys : List (FST i A)} + lstep-there : ∀ {i} {x y : FSTA i} {xs ys : List (FSTA i)} → x ≉ y → (x ∷ xs) lies-in ys ------------------------- → (x ∷ xs) lies-in (y ∷ ys) - _slice-of_ : ∀ {i} → FST i A → FST i A → Set + _slice-of_ : ∀ {i} → FSTA i → FSTA i → Set x slice-of y = (x ∷ []) lies-in (y ∷ []) - _slice-within_ : ∀ {i} → FST i A → List (FST i A) → Set + _slice-within_ : ∀ {i} → FSTA i → List (FSTA i) → Set x slice-within ys = (x ∷ []) lies-in ys - lies-in-refl : ∀ {i} → (xs : List (FST i A)) → xs lies-in xs + lies-in-refl : ∀ {i} → (xs : List (FSTA i)) → xs lies-in xs lies-in-refl [] = lempty lies-in-refl ((a -< as >-) ∷ xs) = lstep-here refl (lies-in-refl as) (lies-in-refl xs) - slice-prop : ∀ {i} {xs ys : List (FST i A)} (zs : List (FST i A)) + slice-prop : ∀ {i} {xs ys : List (FSTA i)} (zs : List (FSTA i)) → xs lies-in ys → xs lies-in (ys ⊕ zs) slice-prop zs lempty = lempty slice-prop {xs = a -< as >- ∷ xs} {ys = .a -< bs >- ∷ ys} zs (lstep-here refl as-lies-in-bs xs-lies-in-ys) = {!lstep-here!} slice-prop zs (lstep-there x x₁) = {!!} - slice-concat : ∀ {i} {x : FST i A} {ys : List (FST i A)} (xs : List (FST i A)) + slice-concat : ∀ {i} {x : FSTA i} {ys : List (FSTA i)} (xs : List (FSTA i)) → x slice-within ys → (x ∷ xs) lies-in (ys ⊕ xs) slice-concat = {!!} -- mutual - -- ⊕-makes-slicesˡ : ∀ {i} (xs ys : List (FST i A)) + -- ⊕-makes-slicesˡ : ∀ {i} (xs ys : List (FSTA i)) -- → xs lies-in (ys ⊕ xs) -- ⊕-makes-slicesˡ [] ys = lempty -- ⊕-makes-slicesˡ (x ∷ xs) ys = slice-concat xs (⊙-makes-slice-head x ys) - -- ⊕-makes-slicesʳ : ∀ {i} (xs ys : List (FST i A)) + -- ⊕-makes-slicesʳ : ∀ {i} (xs ys : List (FSTA i)) -- → xs lies-in (xs ⊕ ys) -- ⊕-makes-slicesʳ xs [] = lies-in-refl xs -- ⊕-makes-slicesʳ xs (y ∷ ys) = slice-prop ys (⊙-makes-slice-tail y xs) - -- ⊙-makes-slice-tail : ∀ {i} (x : FST i A) (ys : List (FST i A)) + -- ⊙-makes-slice-tail : ∀ {i} (x : FSTA i) (ys : List (FSTA i)) -- → ys lies-in (x ⊙ ys) -- ⊙-makes-slice-tail x [] = lempty -- ⊙-makes-slice-tail (a -< cs >-) ((b -< bs >-) ∷ ys) with a ≟ b -- ... | yes refl = lstep-here refl (⊕-makes-slicesˡ bs cs) (lies-in-refl ys) -- ... | no _ = lstep-here refl (lies-in-refl bs) (⊙-makes-slice-tail (a -< cs >-) ys) - -- ⊙-makes-slice-head : ∀ {i} (x : FST i A) (ys : List (FST i A)) + -- ⊙-makes-slice-head : ∀ {i} (x : FSTA i) (ys : List (FSTA i)) -- → x slice-within (x ⊙ ys) -- ⊙-makes-slice-head (a -< cs >-) [] = lies-in-refl (a -< cs >- ∷ []) -- ⊙-makes-slice-head (a -< cs >-) ((b -< bs >-) ∷ ys) with a ≟ b -- ... | yes refl = lstep-here refl (⊕-makes-slicesʳ cs bs) lempty -- ... | no a≠b = lstep-there a≠b (⊙-makes-slice-head (a -< cs >-) ys) - ⊕-idem : ∀ {i} (xs ys : List (FST i A)) + ⊕-idem : ∀ {i} (xs ys : List (FSTA i)) → AllWellFormed xs → AllWellFormed ys → ys ⊕ xs ⊕ ys ≡ xs ⊕ ys @@ -338,11 +349,11 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where record FSF : Set where constructor _⊚_ field - trees : List (FST ∞ A) + trees : List (FSTA ∞) valid : AllWellFormed trees open FSF public - forget-uniqueness : FSF → List (FST ∞ A) + forget-uniqueness : FSF → List (FSTA ∞) forget-uniqueness = trees {- @@ -438,7 +449,7 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where open import Data.Product using (_,_) -- Semantics - ⟦_⟧ : SPL → Conf → Rose ∞ A + ⟦_⟧ : SPL → Conf → Rose ∞ AtomSet ⟦ r ◀ features ⟧ c = r -< forget-uniqueness (⊛-all (select c features)) >- open import Data.String using (String; _<+>_) @@ -446,12 +457,12 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where module Show (show-F : F → String) (show-A : A → String) where mutual - show-FST : FST ∞ A → Lines + show-FST : FSTA ∞ → Lines show-FST = induction λ a children → do > show-A a indent 2 (lines children) - show-FSF : List (FST ∞ A) → Lines + show-FSF : List (FSTA ∞) → Lines show-FSF roots = lines (map show-FST roots) show-Feature : Feature → Lines @@ -459,20 +470,8 @@ module Impose {A : 𝔸} (_≟_ : DecidableEquality A) where > show-F (name feature) <+> "∷" indent 2 (show-FSF (forget-uniqueness (impl feature))) --- Our framework does not allow constraints on the atom type A. --- This demonstrates that FSTs are not totally generic. --- DecidableEquality is a reasonable constraint though - or rather it is an axiom. --- Maybe we could avoid this constraint by not comparing atoms A --- but equipping each node with an additional ID and compare that instead? --- That would allow merging nodes with unequal atoms though. --- Maybe assuming decidable equality as an axiom here is just fine? --- Not having decidable equality is far from practical. -module Framework (mkDec : (A : 𝔸) → DecidableEquality A) where - FSTL-𝔼 : 𝔼 - FSTL-𝔼 A = Impose.SPL (mkDec A) - - FSTL-Sem : 𝔼-Semantics (Rose ∞) Conf FSTL-𝔼 - FSTL-Sem {A} = Impose.⟦_⟧ (mkDec A) - - FSTL : VariabilityLanguage (Rose ∞) - FSTL = ⟪ FSTL-𝔼 , Conf , FSTL-Sem ⟫ +FSTL-Sem : 𝔼-Semantics (Rose ∞) Conf Impose.SPL +FSTL-Sem {A} = Impose.⟦_⟧ A + +FSTL : VariabilityLanguage (Rose ∞) +FSTL = ⟪ Impose.SPL , Conf , FSTL-Sem ⟫ From 8a03b95b41ad73325b0600b7bffaf3ba76b74a4e Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 5 Apr 2024 22:30:49 +0200 Subject: [PATCH 42/43] repair FST experiment --- src/Lang/FST.agda | 9 +++- src/Test/Experiments/FST-Experiments.agda | 65 ++++++++++++----------- 2 files changed, 41 insertions(+), 33 deletions(-) diff --git a/src/Lang/FST.agda b/src/Lang/FST.agda index b9d9857b..a75925df 100644 --- a/src/Lang/FST.agda +++ b/src/Lang/FST.agda @@ -160,8 +160,10 @@ branches = concat module Impose (AtomSet : 𝔸) where FSTA : Size → Set FSTA i = FST i AtomSet - A = atoms AtomSet - _≟_ = proj₂ AtomSet + + private + A = atoms AtomSet + _≟_ = proj₂ AtomSet _==_ : ∀ {i} → Decidable (_≈_ {AtomSet} {i}) (a -< _ >-) == (b -< _ >-) = a ≟ b @@ -388,6 +390,9 @@ module Impose (AtomSet : 𝔸) where then impl f ∷ select c fs else select c fs + names : SPL → List (Name F) + names spl = (map name) (features spl) + ---- Algebra ---- 𝟘 : FSF diff --git a/src/Test/Experiments/FST-Experiments.agda b/src/Test/Experiments/FST-Experiments.agda index 38beb57f..285acb56 100644 --- a/src/Test/Experiments/FST-Experiments.agda +++ b/src/Test/Experiments/FST-Experiments.agda @@ -7,7 +7,7 @@ module Test.Experiments.FST-Experiments where open import Data.Bool using (true; false) open import Data.List using (List; _∷_; []; map; [_]) open import Data.Product using (proj₁; proj₂; _,_; _×_) -open import Function using (id) +open import Function using (id; _∘_) open import Relation.Binary using (DecidableEquality) open import Relation.Nullary.Decidable using (yes; no; does; _because_; _×-dec_) @@ -21,33 +21,36 @@ open import Show.Lines open import Util.ShowHelpers open import Data.String using (String; _<+>_; _++_) renaming (_≟_ to _≟ˢ_) -open import Lang.FST as FSTModule using (Conf) +open import Framework.Variants using (show-rose) -module _ {A : 𝔸} (_≟_ : DecidableEquality A) where - open FSTModule.Defs {A} - open FSTModule.Defs.Impose _≟_ - module FSTShow = FSTModule.Defs.Impose.Show +import Lang.FST as FST +open FST using (Conf) - exp : ∀ {N} - → (N → String) - → (A → String) - → List (Conf N) - → Experiment (SPL N) +module _ (F : 𝔽) (A : 𝔸) where +-- (_≟_ : DecidableEquality A) + open FST.Impose F A + module FSTShow = FST.Impose.Show F A + + exp : + (F → String) + → (atoms A → String) + → List (Conf F) + → Experiment SPL getName (exp _ _ _) = "Configure FST example" get (exp show-N show-A configs) (example-name ≔ forest) = - let open FSTShow _≟_ show-N show-A + let open FSTShow show-N show-A in do > "Expression e has features" indent 2 do - lines (map show-Feature forest) + lines (map show-Feature (features forest)) foreach [ c ∈ configs ] do let cstr = show-fun show-N show-bool c (names forest) linebreak > "⟦ e ⟧" <+> cstr <+> "=" indent 2 do - show-FSF (forget-uniqueness (⟦ forest ⟧ c)) + > show-rose show-A (⟦ forest ⟧ c) pick-all : ∀ {N} → Conf N pick-all _ = true @@ -66,8 +69,8 @@ module Java where _≟-ast_ : DecidableEquality ASTNode _≟-ast_ = _≟ˢ_ - open FSTModule.Defs {ASTNode} - open FSTModule.Defs.Impose _≟-ast_ + open FST String using (_._; branches) + open FST.Impose String (ASTNode , _≟-ast_) module Calculator where fname-Add = "Add" @@ -90,35 +93,35 @@ module Java where open import Data.List.Relation.Unary.AllPairs using ([]; _∷_) open import Data.List.Relation.Unary.All using ([]; _∷_) - feature-Add : Feature ASTNode - feature-Add = fname-Add :: (cls . add . add-ret . [] , ([] ∷ []) , ((unq ([] ∷ [] , (unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ [])) ∷ [])) + feature-Add : Feature + feature-Add = fname-Add :: (cls . add . add-ret . []) ⊚ (([] ∷ []) , ((([] ∷ []) , ((([] ∷ []) , (([] , []) ∷ [])) ∷ [])) ∷ [])) - feature-Sub : Feature ASTNode - feature-Sub = fname-Sub :: (cls . sub . sub-ret . [] , ([] ∷ []) , ((unq ([] ∷ [] , (unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ [])) ∷ [])) + feature-Sub : Feature + feature-Sub = fname-Sub :: (cls . sub . sub-ret . []) ⊚ (([] ∷ []) , ((([] ∷ [] , ((([] ∷ []) , ((([] , [])) ∷ []))) ∷ [])) ∷ [])) - feature-Log : Feature ASTNode - feature-Log = fname-Log :: cls . + feature-Log : Feature + feature-Log = fname-Log :: (cls . branches ( (add . log . []) ∷ (sub . log . []) - ∷ []) , ([] ∷ []) , ((unq (((tt ∷ []) ∷ ([] ∷ [])) , ((unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ ((unq (([] ∷ []) , ((unq ([] , [])) ∷ []))) ∷ [])))) ∷ []) + ∷ [])) ⊚ (([] ∷ []) , ((((((λ where ()) ∷ []) ∷ ([] ∷ [])) , (((([] ∷ []) , ((([] , [])) ∷ []))) ∷ (((([] ∷ []) , ((([] , [])) ∷ []))) ∷ [])))) ∷ [])) ---- Example SPLs - ex-Add-Sub : Example (SPL ASTNode) - ex-Add-Sub = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ [] + ex-Add-Sub : Example SPL + ex-Add-Sub = "add-sub" ≔ "package" ◀ (feature-Add ∷ feature-Sub ∷ []) - ex-Sub-Add : Example (SPL ASTNode) - ex-Sub-Add = "sub-add" ≔ feature-Sub ∷ feature-Add ∷ [] + ex-Sub-Add : Example SPL + ex-Sub-Add = "sub-add" ≔ "package" ◀ (feature-Sub ∷ feature-Add ∷ []) - ex-Add-Sub-Log : Example (SPL ASTNode) - ex-Add-Sub-Log = "add-sub" ≔ feature-Add ∷ feature-Sub ∷ feature-Log ∷ [] + ex-Add-Sub-Log : Example SPL + ex-Add-Sub-Log = "add-sub" ≔ "package" ◀ (feature-Add ∷ feature-Sub ∷ feature-Log ∷ []) - ex-all : List (Example (SPL ASTNode)) + ex-all : List (Example SPL) ex-all = ex-Add-Sub ∷ ex-Sub-Add ∷ ex-Add-Sub-Log ∷ [] ---- Experiments toy-calculator-experiment = let eq = _≟-ast_ in - exp eq id id (pick-all ∷ pick-only eq fname-Add ∷ []) + exp String (ASTNode , eq) id id (pick-all ∷ pick-only eq fname-Add ∷ []) From 12167af880766c7e919378c9d4811200cbc90d4a Mon Sep 17 00:00:00 2001 From: pmbittner Date: Fri, 5 Apr 2024 22:42:55 +0200 Subject: [PATCH 43/43] repair FST-to-OC --- src/Translation/Lang/FST-to-OC.agda | 264 +++++++++++----------------- 1 file changed, 99 insertions(+), 165 deletions(-) diff --git a/src/Translation/Lang/FST-to-OC.agda b/src/Translation/Lang/FST-to-OC.agda index c3cfe406..0ed1cf40 100644 --- a/src/Translation/Lang/FST-to-OC.agda +++ b/src/Translation/Lang/FST-to-OC.agda @@ -5,7 +5,9 @@ module Translation.Lang.FST-to-OC (F : 𝔽) where open import Size using (Size; ↑_; _⊔ˢ_; ∞) +open import Data.Nat using (_≟_) open import Data.List using (List; []; _∷_; map) +open import Data.Product using (_,_) open import Relation.Nullary.Decidable using (yes; no; _because_; False) open import Relation.Binary using (Decidable; DecidableEquality; Rel) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl) @@ -16,171 +18,103 @@ V = Rose ∞ mkArtifact = Artifact∈ₛRose Option = F +open import Framework.Relation.Expressiveness V + open import Framework.VariabilityLanguage open import Construct.Artifact as At using () import Lang.FST as FST --- open FST F using (FST; induction; fst-leaf) - -open import Framework.Relation.Expressiveness V - --- module _ {A : 𝔸} (_==_ : DecidableEquality A) where --- open FST.Impose F _==_ using (SPL; _◀_; Feature; _::_; FSF; _⊚_) - --- -- translate-o : ∀ {i} → A → OC i A → Feature --- -- translate-o a (b OC.-< cs >-) = {!!} --- -- translate-o a (O OC.❲ e ❳) = {!!} - --- -- left is base --- mutual --- embed : FST A → OC ∞ A --- embed = induction OC._-<_>- - --- merge-trees : F → List (OC ∞ A) → F → FST A → List (OC ∞ A) --- merge-trees _ [] P t = P ❲ embed t ❳ ∷ [] --- merge-trees O (a OC.-< as >- ∷ inter) P (b FST.-< bs >-) with a == b --- ... | yes refl = a OC.-< merge O as P bs >- ∷ inter --- ... | no nope = O ❲ a OC.-< as >- ❳ ∷ merge-trees O inter P (b FST.-< bs >-) --- merge-trees O (Q ❲ e ❳ ∷ inter) P t = Q ❲ e ❳ ∷ merge-trees O inter P t - --- {-# TERMINATING #-} --- merge : F → List (OC ∞ A) → F → List (FST A) → List (OC ∞ A) --- merge _ ls _ [] = ls --- merge O ls P (t ∷ ts) = merge O (merge-trees O ls P t) P ts - --- record Intermediate : Set where --- constructor _:o:_ --- field --- name : F --- children : List (OC ∞ A) - --- translate' : List Intermediate → List (OC ∞ A) --- translate' [] = [] --- translate' ((_ :o: cs) ∷ []) = cs --- translate' ((O :o: os) ∷ (P :o: ps) ∷ xs) = translate' ({!merge!} ∷ xs) - --- fromFeature : Feature → Intermediate --- fromFeature (O :: (ts ⊚ _)) = O :o: map embed ts - --- start : List Feature → List Intermediate --- start = map fromFeature - --- translate : SPL → WFOC ∞ A --- translate (a ◀ fs) = Root a (translate' (start fs)) - -module _ (mkDec : (A : 𝔸) → DecidableEquality A) where - data FeatureName : Set where - X : FeatureName - Y : FeatureName - - open import Data.Bool using (true; false; if_then_else_) - open import Data.Nat using (zero; suc; ℕ) - open import Data.Fin using (Fin; zero; suc) - open import Data.List.Relation.Unary.All using (All; []; _∷_) - open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) - open import Framework.VariantMap V ℕ - - import Lang.OC - open Lang.OC FeatureName using (OC; _❲_❳; WFOC; Root; opt; Configuration) - open Lang.OC.Sem FeatureName V mkArtifact - - open FST FeatureName using (_._) - open FST.Impose FeatureName (mkDec ℕ) using (SPL; _◀_; _::_; _⊚_; unq) renaming (⟦_⟧ to FST⟦_⟧) - open FST.Framework FeatureName mkDec - - open import Data.EqIndexedSet - open import Data.Empty using (⊥-elim) - - open import Data.Product using (_,_; ∃-syntax) - open import Util.Existence using (¬Σ-syntax) - - counterexample : VMap 3 - counterexample zero = rose-leaf 0 - counterexample (suc zero) = rose (0 At.-< rose (1 At.-< rose-leaf 2 ∷ [] >-) ∷ [] >-) - counterexample (suc (suc zero)) = rose (0 At.-< rose (1 At.-< rose-leaf 3 ∷ [] >-) ∷ [] >-) - counterexample (suc (suc (suc zero))) = rose (0 At.-< rose (1 At.-< rose-leaf 2 ∷ rose-leaf 3 ∷ [] >-) ∷ [] >-) - - -- illegal' : ∀ {i : Size} → ∄[ cs ∈ List (OC i ℕ) ] (⟦ Root zero cs ⟧ ≅ counterexample) - -- illegal' ([] , fst , snd) with snd (suc zero) - -- ... | () - -- illegal' (x ∷ fst , snd) = {!!} - - -- illegal' : ∀ {i : Size} → ∄[ e ∈ OC i ℕ ] (∃[ O ] (∃[ xs ] (⟦ Root zero (opt O e ∷ xs) ⟧ ≅ counterexample))) - -- illegal' x = {!!} - open import Relation.Nullary.Negation using (¬_) - - nodup : ∀ {i : Size} {A : 𝔸} (a : A) (x y : OC i A) (zs : List (OC i A)) - → ¬ (∀ (c : Configuration) → (⟦ Root a (x ∷ y ∷ zs) ⟧ c ≡ rose-leaf a)) - nodup a x y zs sure with sure (λ _ → true) - nodup a (O ❲ e ❳) y zs sure | asd = {!!} - -- nodup a (b OC.-< bs >-) y zs sure with sure (λ _ → true) - -- ... | () - -- nodup a (O ❲ e ❳) (b OC.-< bs >-) zs sure with sure (λ _ → false) - -- ... | () - -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] sure = {!!} - -- nodup a (O ❲ e ❳) (P ❲ f ❳) (z ∷ zs) sure = nodup a (P ❲ f ❳) z zs {!sure!} - -- nodup a (O ❲ e ❳) y zs (c , sure) with c O - -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] (c , sure) | false with c P - -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] (c , refl) | false | false = {!!} - -- nodup a (O ❲ e ❳) (P ❲ f ❳) [] (c , sure) | false | true = {!!} - -- nodup a (O ❲ e ❳) (P ❲ f ❳) (z ∷ zs) (c , sure) | false = nodup a (P ❲ f ❳) z zs (c , sure) - -- nodup a (O ❲ e ❳) y zs (c , sure) | true = {!!} - - illegal : ∀ {i : Size} → ∄[ e ∈ WFOC i ℕ ] (⟦ e ⟧ ≅ counterexample) - -- root must be zero because it is always zero in counterexample - illegal (Root (suc i) cs , _ , ⊇) with ⊇ zero - ... | () - -- there must be a child because there are variants in counterexample with children below the root (e.g., suc zero) - illegal (Root zero [] , _ , ⊇) with ⊇ (suc zero) -- illegal' (cs , eq) - ... | () - -- there must be an option at the front because there are variants with zero children (e.g., zero) - illegal (Root zero (a OC.-< cs >- ∷ _) , _ , ⊇) with ⊇ zero - ... | () - -- there can not be any other children - illegal (Root zero (O ❲ e ❳ ∷ a OC.-< as >- ∷ xs) , ⊆ , ⊇) = {!!} - illegal (Root zero (O ❲ e ❳ ∷ P OC.❲ e' ❳ ∷ xs) , ⊆ , ⊇) = {!!} - -- e can be a chain of options but somewhen, there must be an artifact - illegal (Root zero (O ❲ e ❳ ∷ []) , ⊆ , ⊇) = {!!} - --illegal' (e , (O , xs , (⊆ , ⊇))) - - cef : SPL - cef = 0 ◀ ( - (X :: (1 . 2 . []) ⊚ ([] ∷ [] , unq ([] ∷ [] , unq ([] , []) ∷ []) ∷ [])) ∷ - (Y :: (1 . 3 . []) ⊚ ([] ∷ [] , unq ([] ∷ [] , unq ([] , []) ∷ []) ∷ [])) ∷ - []) - - cef-describes-counterexample : FST⟦ cef ⟧ ≅ counterexample - cef-describes-counterexample = ⊆[]→⊆ cef-⊆[] , ⊆[]→⊆ {f = fnoc} cef-⊇[] - where - conf : FST.Conf FeatureName → Fin 4 - conf c with c X | c Y - ... | false | false = zero - ... | false | true = suc (suc zero) - ... | true | false = suc zero - ... | true | true = suc (suc (suc zero)) - - cef-⊆[] : FST⟦ cef ⟧ ⊆[ conf ] counterexample - cef-⊆[] c with c X | c Y - ... | false | false = refl - ... | false | true = refl - ... | true | false = refl - ... | true | true = {!!} - - fnoc : Fin 4 → FST.Conf FeatureName - fnoc zero X = false - fnoc zero Y = false - fnoc (suc zero) X = true - fnoc (suc zero) Y = false - fnoc (suc (suc zero)) X = false - fnoc (suc (suc zero)) Y = true - fnoc (suc (suc (suc zero))) X = true - fnoc (suc (suc (suc zero))) Y = true - - cef-⊇[] : counterexample ⊆[ fnoc ] FST⟦ cef ⟧ - cef-⊇[] zero = refl - cef-⊇[] (suc zero) = refl - cef-⊇[] (suc (suc zero)) = refl - cef-⊇[] (suc (suc (suc zero))) = {!!} - - ouch : WFOCL ⋡ FSTL - ouch sure with sure cef - ... | e , e≣cef = {!!} +open FST F using (FST; FSTL) + +data FeatureName : Set where + X : FeatureName + Y : FeatureName + +open import Data.Bool using (true; false; if_then_else_) +open import Data.Nat using (zero; suc; ℕ) +open import Data.Fin using (Fin; zero; suc) +open import Data.List.Relation.Unary.All using (All; []; _∷_) +open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Framework.VariantMap V (ℕ , _≟_) + +import Lang.All +open Lang.All.OC using (OC; _❲_❳; WFOC; Root; opt; Configuration; ⟦_⟧; WFOCL) + +AtomSet : 𝔸 +AtomSet = ℕ , _≟_ + +open FST FeatureName using (_._) +open FST.Impose FeatureName AtomSet using (SPL; _◀_; _::_; _⊚_) renaming (⟦_⟧ to FST⟦_⟧) + +open import Data.EqIndexedSet +open import Data.Empty using (⊥-elim) + +open import Data.Product using (_,_; ∃-syntax) +open import Util.Existence using (¬Σ-syntax) + +counterexample : VMap 3 +counterexample zero = rose-leaf 0 +counterexample (suc zero) = rose (0 At.-< rose (1 At.-< rose-leaf 2 ∷ [] >-) ∷ [] >-) +counterexample (suc (suc zero)) = rose (0 At.-< rose (1 At.-< rose-leaf 3 ∷ [] >-) ∷ [] >-) +counterexample (suc (suc (suc zero))) = rose (0 At.-< rose (1 At.-< rose-leaf 2 ∷ rose-leaf 3 ∷ [] >-) ∷ [] >-) + +open import Relation.Nullary.Negation using (¬_) + + +illegal : ∀ {i : Size} → ∄[ e ∈ WFOC FeatureName i AtomSet ] (⟦ e ⟧ ≅ counterexample) +-- root must be zero because it is always zero in counterexample +illegal (Root (suc i) cs , _ , ⊇) with ⊇ zero +... | () +-- there must be a child because there are variants in counterexample with children below the root (e.g., suc zero) +illegal (Root zero [] , _ , ⊇) with ⊇ (suc zero) -- illegal' (cs , eq) +... | () +-- there must be an option at the front because there are variants with zero children (e.g., zero) +illegal (Root zero (a OC.-< cs >- ∷ _) , _ , ⊇) with ⊇ zero +... | () +-- there can not be any other children +illegal (Root zero (O ❲ e ❳ ∷ a OC.-< as >- ∷ xs) , ⊆ , ⊇) = {!!} +illegal (Root zero (O ❲ e ❳ ∷ P OC.❲ e' ❳ ∷ xs) , ⊆ , ⊇) = {!!} +-- e can be a chain of options but somewhen, there must be an artifact +illegal (Root zero (O ❲ e ❳ ∷ []) , ⊆ , ⊇) = {!!} +--illegal' (e , (O , xs , (⊆ , ⊇))) + +cef : SPL +cef = 0 ◀ ( + (X :: (1 . 2 . []) ⊚ ([] ∷ [] , ([] ∷ [] , ([] , []) ∷ []) ∷ [])) ∷ + (Y :: (1 . 3 . []) ⊚ ([] ∷ [] , ([] ∷ [] , ([] , []) ∷ []) ∷ [])) ∷ + []) + +cef-describes-counterexample : FST⟦ cef ⟧ ≅ counterexample +cef-describes-counterexample = ⊆[]→⊆ cef-⊆[] , ⊆[]→⊆ {f = fnoc} cef-⊇[] + where + conf : FST.Conf FeatureName → Fin 4 + conf c with c X | c Y + ... | false | false = zero + ... | false | true = suc (suc zero) + ... | true | false = suc zero + ... | true | true = suc (suc (suc zero)) + + cef-⊆[] : FST⟦ cef ⟧ ⊆[ conf ] counterexample + cef-⊆[] c with c X | c Y + ... | false | false = refl + ... | false | true = refl + ... | true | false = refl + ... | true | true = {!!} + + fnoc : Fin 4 → FST.Conf FeatureName + fnoc zero X = false + fnoc zero Y = false + fnoc (suc zero) X = true + fnoc (suc zero) Y = false + fnoc (suc (suc zero)) X = false + fnoc (suc (suc zero)) Y = true + fnoc (suc (suc (suc zero))) X = true + fnoc (suc (suc (suc zero))) Y = true + + cef-⊇[] : counterexample ⊆[ fnoc ] FST⟦ cef ⟧ + cef-⊇[] zero = refl + cef-⊇[] (suc zero) = refl + cef-⊇[] (suc (suc zero)) = refl + cef-⊇[] (suc (suc (suc zero))) = {!!} + +ouch : WFOCL FeatureName ⋡ FSTL +ouch sure = {!!}