summaryrefslogtreecommitdiff
path: root/src/Cfe/Language/Construct/Union.agda
diff options
context:
space:
mode:
authorChloe Brown <chloe.brown.00@outlook.com>2021-03-05 00:22:20 +0000
committerChloe Brown <chloe.brown.00@outlook.com>2021-03-05 00:22:20 +0000
commit90c7597b6f80e0bd5bb2a9a7245d097d40486518 (patch)
treef106450275fb3375d0346a2f02e72f6537f8c5cc /src/Cfe/Language/Construct/Union.agda
parent9c5540562f53be81bfb03ef2b96d7f59b3ddea44 (diff)
Prove lemma 3.3
Diffstat (limited to 'src/Cfe/Language/Construct/Union.agda')
-rw-r--r--src/Cfe/Language/Construct/Union.agda8
1 files changed, 7 insertions, 1 deletions
diff --git a/src/Cfe/Language/Construct/Union.agda b/src/Cfe/Language/Construct/Union.agda
index 4ed0774..5099d04 100644
--- a/src/Cfe/Language/Construct/Union.agda
+++ b/src/Cfe/Language/Construct/Union.agda
@@ -25,7 +25,7 @@ module _
where
infix 4 _≈ᵁ_
- infix 4 _∪_
+ infix 6 _∪_
Union : List C → Set (a ⊔ b)
Union l = l ∈ A ⊎ l ∈ B
@@ -93,3 +93,9 @@ isCommutativeMonoid = record
where
module X≤Y = _≤_ X≤Y
module U≤V = _≤_ U≤V
+
+∪-unique : ∀ {a aℓ b bℓ} {A : Language a aℓ} {B : Language b bℓ} → (∀ x → first A x → first B x → ⊥) → (𝕃.null A → 𝕃.null B → ⊥) → ∀ {l} → l ∈ A ∪ B → (l ∈ A × l ∉ B) ⊎ (l ∉ A × l ∈ B)
+∪-unique fA∩fB≡∅ ¬nA∨¬nB {[]} (inj₁ []∈A) = inj₁ ([]∈A , ¬nA∨¬nB []∈A)
+∪-unique fA∩fB≡∅ ¬nA∨¬nB {x ∷ l} (inj₁ xl∈A) = inj₁ (xl∈A , (λ xl∈B → fA∩fB≡∅ x (-, xl∈A) (-, xl∈B)))
+∪-unique fA∩fB≡∅ ¬nA∨¬nB {[]} (inj₂ []∈B) = inj₂ (flip ¬nA∨¬nB []∈B , []∈B)
+∪-unique fA∩fB≡∅ ¬nA∨¬nB {x ∷ l} (inj₂ xl∈B) = inj₂ ((λ xl∈A → fA∩fB≡∅ x (-, xl∈A) (-, xl∈B)) , xl∈B)