141
|
1 module regular-concat where
|
|
2
|
|
3 open import Level renaming ( suc to Suc ; zero to Zero )
|
|
4 open import Data.List
|
|
5 open import Data.Nat hiding ( _≟_ )
|
|
6 open import Data.Fin hiding ( _+_ )
|
|
7 open import Data.Empty
|
|
8 open import Data.Unit
|
|
9 open import Data.Product
|
|
10 -- open import Data.Maybe
|
|
11 open import Relation.Nullary
|
|
12 open import Relation.Binary.PropositionalEquality hiding ( [_] )
|
|
13 open import logic
|
|
14 open import nat
|
|
15 open import automaton
|
|
16 open import regular-language
|
|
17
|
|
18 open import nfa
|
|
19 open import sbconst2
|
|
20
|
|
21 open RegularLanguage
|
|
22 open Automaton
|
|
23
|
|
24 Concat-NFA : {Σ : Set} → (A B : RegularLanguage Σ ) → ((x y : states A )→ Dec (x ≡ y)) → ((x y : states B )→ Dec (x ≡ y))
|
|
25 → NAutomaton (states A ∨ states B) Σ
|
|
26 Concat-NFA {Σ} A B equal?A equal?B = record { Nδ = δnfa ; Nend = nend }
|
|
27 module Concat-NFA where
|
|
28 δnfa : states A ∨ states B → Σ → states A ∨ states B → Bool
|
|
29 δnfa (case1 q) i (case1 q₁) with equal?A (δ (automaton A) q i) q₁
|
|
30 ... | yes _ = true
|
|
31 ... | no _ = false
|
|
32 δnfa (case1 qa) i (case2 qb) with equal?B qb (δ (automaton B) (astart B) i)
|
|
33 ... | yes _ = aend (automaton A) qa
|
|
34 ... | no _ = false
|
|
35 δnfa (case2 q) i (case2 q₁) with equal?B (δ (automaton B) q i) q₁
|
|
36 ... | yes _ = true
|
|
37 ... | no _ = false
|
|
38 δnfa _ i _ = false
|
|
39 nend : states A ∨ states B → Bool
|
|
40 nend (case2 q) = aend (automaton B) q
|
|
41 nend (case1 q) = aend (automaton A) q /\ aend (automaton B) (astart B) -- empty B case
|
|
42
|
|
43 Concat-NFA-start : {Σ : Set} → (A B : RegularLanguage Σ ) → states A ∨ states B → ((x y : states A )→ Dec (x ≡ y)) → Bool
|
|
44 Concat-NFA-start A B (case1 a) equal?A with equal?A a (astart A)
|
|
45 ... | yes _ = true
|
|
46 ... | no _ = false
|
|
47 Concat-NFA-start A B (case2 b) equal?A = false
|
|
48
|
|
49 M-Concat : {Σ : Set} → (A B : RegularLanguage Σ ) → ((states A → Bool) → Bool) → ((states B → Bool) → Bool) → RegularLanguage Σ
|
|
50 M-Concat {Σ} A B existsA existsB = record {
|
|
51 states = states A ∨ states B → Bool
|
|
52 ; astart = λ ab → Concat-NFA-start A B ab {!!}
|
|
53 ; automaton = subset-construction sbexists (Concat-NFA A B {!!} {!!} )
|
|
54 } where
|
|
55 sbexists : (states A ∨ states B → Bool) → Bool
|
|
56 sbexists P = existsA ( λ a → existsB ( λ b → P (case1 a) \/ P (case2 b)))
|
|
57
|
|
58 record Split {Σ : Set} (A : List Σ → Bool ) ( B : List Σ → Bool ) (x : List Σ ) : Set where
|
|
59 field
|
|
60 sp0 : List Σ
|
|
61 sp1 : List Σ
|
|
62 sp-concat : sp0 ++ sp1 ≡ x
|
|
63 prop0 : A sp0 ≡ true
|
|
64 prop1 : B sp1 ≡ true
|
|
65
|
|
66 open Split
|
|
67
|
|
68 list-empty++ : {Σ : Set} → (x y : List Σ) → x ++ y ≡ [] → (x ≡ [] ) ∧ (y ≡ [] )
|
|
69 list-empty++ [] [] refl = record { proj1 = refl ; proj2 = refl }
|
|
70 list-empty++ [] (x ∷ y) ()
|
|
71 list-empty++ (x ∷ x₁) y ()
|
|
72
|
|
73 open _∧_
|
|
74
|
|
75 open import Relation.Binary.PropositionalEquality hiding ( [_] )
|
|
76
|
|
77 c-split-lemma : {Σ : Set} → (A B : List Σ → Bool ) → (h : Σ) → ( t : List Σ ) → split A B (h ∷ t ) ≡ true
|
|
78 → ( ¬ (A [] ≡ true )) ∨ ( ¬ (B ( h ∷ t ) ≡ true) )
|
|
79 → split (λ t1 → A (h ∷ t1)) B t ≡ true
|
|
80 c-split-lemma {Σ} A B h t eq p = sym ( begin
|
|
81 true
|
|
82 ≡⟨ sym eq ⟩
|
|
83 split A B (h ∷ t )
|
|
84 ≡⟨⟩
|
|
85 A [] /\ B (h ∷ t) \/ split (λ t1 → A (h ∷ t1)) B t
|
|
86 ≡⟨ cong ( λ k → k \/ split (λ t1 → A (h ∷ t1)) B t ) (lemma-p p ) ⟩
|
|
87 false \/ split (λ t1 → A (h ∷ t1)) B t
|
|
88 ≡⟨ bool-or-1 refl ⟩
|
|
89 split (λ t1 → A (h ∷ t1)) B t
|
|
90 ∎ ) where
|
|
91 open ≡-Reasoning
|
|
92 lemma-p : ( ¬ (A [] ≡ true )) ∨ ( ¬ (B ( h ∷ t ) ≡ true) ) → A [] /\ B (h ∷ t) ≡ false
|
|
93 lemma-p (case1 ¬A ) = bool-and-1 ( ¬-bool-t ¬A )
|
|
94 lemma-p (case2 ¬B ) = bool-and-2 ( ¬-bool-t ¬B )
|
|
95
|
|
96 split→AB : {Σ : Set} → (A B : List Σ → Bool ) → ( x : List Σ ) → split A B x ≡ true → Split A B x
|
|
97 split→AB {Σ} A B [] eq with bool-≡-? (A []) true | bool-≡-? (B []) true
|
|
98 split→AB {Σ} A B [] eq | yes eqa | yes eqb =
|
|
99 record { sp0 = [] ; sp1 = [] ; sp-concat = refl ; prop0 = eqa ; prop1 = eqb }
|
|
100 split→AB {Σ} A B [] eq | yes p | no ¬p = ⊥-elim (lemma-∧-1 eq (¬-bool-t ¬p ))
|
|
101 split→AB {Σ} A B [] eq | no ¬p | t = ⊥-elim (lemma-∧-0 eq (¬-bool-t ¬p ))
|
|
102 split→AB {Σ} A B (h ∷ t ) eq with bool-≡-? (A []) true | bool-≡-? (B (h ∷ t )) true
|
|
103 ... | yes px | yes py = record { sp0 = [] ; sp1 = h ∷ t ; sp-concat = refl ; prop0 = px ; prop1 = py }
|
|
104 ... | no px | _ with split→AB (λ t1 → A ( h ∷ t1 )) B t (c-split-lemma A B h t eq (case1 px) )
|
|
105 ... | S = record { sp0 = h ∷ sp0 S ; sp1 = sp1 S ; sp-concat = cong ( λ k → h ∷ k) (sp-concat S) ; prop0 = prop0 S ; prop1 = prop1 S }
|
|
106 split→AB {Σ} A B (h ∷ t ) eq | _ | no px with split→AB (λ t1 → A ( h ∷ t1 )) B t (c-split-lemma A B h t eq (case2 px) )
|
|
107 ... | S = record { sp0 = h ∷ sp0 S ; sp1 = sp1 S ; sp-concat = cong ( λ k → h ∷ k) (sp-concat S) ; prop0 = prop0 S ; prop1 = prop1 S }
|
|
108
|
|
109 AB→split : {Σ : Set} → (A B : List Σ → Bool ) → ( x y : List Σ ) → A x ≡ true → B y ≡ true → split A B (x ++ y ) ≡ true
|
|
110 AB→split {Σ} A B [] [] eqa eqb = begin
|
|
111 split A B []
|
|
112 ≡⟨⟩
|
|
113 A [] /\ B []
|
|
114 ≡⟨ cong₂ (λ j k → j /\ k ) eqa eqb ⟩
|
|
115 true
|
|
116 ∎ where open ≡-Reasoning
|
|
117 AB→split {Σ} A B [] (h ∷ y ) eqa eqb = begin
|
|
118 split A B (h ∷ y )
|
|
119 ≡⟨⟩
|
|
120 A [] /\ B (h ∷ y) \/ split (λ t1 → A (h ∷ t1)) B y
|
|
121 ≡⟨ cong₂ (λ j k → j /\ k \/ split (λ t1 → A (h ∷ t1)) B y ) eqa eqb ⟩
|
|
122 true /\ true \/ split (λ t1 → A (h ∷ t1)) B y
|
|
123 ≡⟨⟩
|
|
124 true \/ split (λ t1 → A (h ∷ t1)) B y
|
|
125 ≡⟨⟩
|
|
126 true
|
|
127 ∎ where open ≡-Reasoning
|
|
128 AB→split {Σ} A B (h ∷ t) y eqa eqb = begin
|
|
129 split A B ((h ∷ t) ++ y)
|
|
130 ≡⟨⟩
|
|
131 A [] /\ B (h ∷ t ++ y) \/ split (λ t1 → A (h ∷ t1)) B (t ++ y)
|
|
132 ≡⟨ cong ( λ k → A [] /\ B (h ∷ t ++ y) \/ k ) (AB→split {Σ} (λ t1 → A (h ∷ t1)) B t y eqa eqb ) ⟩
|
|
133 A [] /\ B (h ∷ t ++ y) \/ true
|
|
134 ≡⟨ bool-or-3 ⟩
|
|
135 true
|
|
136 ∎ where open ≡-Reasoning
|
|
137
|
|
138 open NAutomaton
|
|
139 open import Data.List.Properties
|
|
140
|
163
|
141 open import finiteSet
|
|
142 open import finiteSetUtil
|
|
143
|
|
144 open FiniteSet
|
|
145
|
|
146 closed-in-concat : {Σ : Set} → (A B : RegularLanguage Σ ) → ( x : List Σ ) → isRegular (Concat (contain A) (contain B)) x ( M-Concat A B )
|
141
|
147 closed-in-concat {Σ} A B x = ≡-Bool-func closed-in-concat→ closed-in-concat← where
|
163
|
148 afin : (A : RegularLanguage Σ ) → FiniteSet A
|
|
149 afin = ?
|
|
150 finab = (fin-∨ (afin A) (afin B))
|
|
151 NFA = (Concat-NFA A B)
|
141
|
152 abmove : (q : states A ∨ states B) → (h : Σ ) → states A ∨ states B
|
|
153 abmove (case1 q) h = case1 (δ (automaton A) q h)
|
|
154 abmove (case2 q) h = case2 (δ (automaton B) q h)
|
|
155 lemma-nmove-ab : (q : states A ∨ states B) → (h : Σ ) → Nδ NFA q h (abmove q h) ≡ true
|
178
|
156 lemma-nmove-ab (case1 q) _ = ? -- equal?-refl (afin A)
|
|
157 lemma-nmove-ab (case2 q) _ = ? -- equal?-refl (afin B)
|
141
|
158 nmove : (q : states A ∨ states B) (nq : states A ∨ states B → Bool ) → (nq q ≡ true) → ( h : Σ ) →
|
163
|
159 exists finab (λ qn → nq qn /\ Nδ NFA qn h (abmove q h)) ≡ true
|
|
160 nmove (case1 q) nq nqt h = found finab (case1 q) ( bool-and-tt nqt (lemma-nmove-ab (case1 q) h) )
|
|
161 nmove (case2 q) nq nqt h = found finab (case2 q) ( bool-and-tt nqt (lemma-nmove-ab (case2 q) h) )
|
141
|
162 acceptB : (z : List Σ) → (q : states B) → (nq : states A ∨ states B → Bool ) → (nq (case2 q) ≡ true) → ( accept (automaton B) q z ≡ true )
|
|
163 → Naccept NFA finab nq z ≡ true
|
|
164 acceptB [] q nq nqt fb = lemma8 where
|
163
|
165 lemma8 : exists finab ( λ q → nq q /\ Nend NFA q ) ≡ true
|
|
166 lemma8 = found finab (case2 q) ( bool-and-tt nqt fb )
|
141
|
167 acceptB (h ∷ t ) q nq nq=q fb = acceptB t (δ (automaton B) q h) (Nmoves NFA finab nq h) (nmove (case2 q) nq nq=q h) fb
|
|
168
|
|
169 acceptA : (y z : List Σ) → (q : states A) → (nq : states A ∨ states B → Bool ) → (nq (case1 q) ≡ true)
|
|
170 → ( accept (automaton A) q y ≡ true ) → ( accept (automaton B) (astart B) z ≡ true )
|
|
171 → Naccept NFA finab nq (y ++ z) ≡ true
|
163
|
172 acceptA [] [] q nq nqt fa fb = found finab (case1 q) (bool-and-tt nqt (bool-and-tt fa fb ))
|
141
|
173 acceptA [] (h ∷ z) q nq nq=q fa fb = acceptB z nextb (Nmoves NFA finab nq h) lemma70 fb where
|
|
174 nextb : states B
|
|
175 nextb = δ (automaton B) (astart B) h
|
163
|
176 lemma70 : exists finab (λ qn → nq qn /\ Nδ NFA qn h (case2 nextb)) ≡ true
|
|
177 lemma70 = found finab (case1 q) ( bool-and-tt nq=q (bool-and-tt fa (lemma-nmove-ab (case2 (astart B)) h) ))
|
141
|
178 acceptA (h ∷ t) z q nq nq=q fa fb = acceptA t z (δ (automaton A) q h) (Nmoves NFA finab nq h) (nmove (case1 q) nq nq=q h) fa fb where
|
|
179
|
|
180 acceptAB : Split (contain A) (contain B) x
|
163
|
181 → Naccept NFA finab (equal? finab (case1 (astart A))) x ≡ true
|
|
182 acceptAB S = subst ( λ k → Naccept NFA finab (equal? finab (case1 (astart A))) k ≡ true ) ( sp-concat S )
|
178
|
183 (acceptA (sp0 S) (sp1 S) (astart A) (equal? finab (case1 (astart A))) ? (prop0 S) (prop1 S) )
|
141
|
184
|
163
|
185 closed-in-concat→ : Concat (contain A) (contain B) x ≡ true → contain (M-Concat A B) x ≡ true
|
|
186 closed-in-concat→ concat with split→AB (contain A) (contain B) x concat
|
|
187 ... | S = begin
|
|
188 accept (subset-construction finab NFA (case1 (astart A))) (Concat-NFA-start A B ) x
|
|
189 ≡⟨ ≡-Bool-func (subset-construction-lemma← finab NFA (case1 (astart A)) x )
|
|
190 (subset-construction-lemma→ finab NFA (case1 (astart A)) x ) ⟩
|
|
191 Naccept NFA finab (equal? finab (case1 (astart A))) x
|
|
192 ≡⟨ acceptAB S ⟩
|
|
193 true
|
|
194 ∎ where open ≡-Reasoning
|
|
195
|
|
196 open Found
|
141
|
197
|
|
198 ab-case : (q : states A ∨ states B ) → (qa : states A ) → (x : List Σ ) → Set
|
|
199 ab-case (case1 qa') qa x = qa' ≡ qa
|
|
200 ab-case (case2 qb) qa x = ¬ ( accept (automaton B) qb x ≡ true )
|
|
201
|
|
202 contain-A : (x : List Σ) → (nq : states A ∨ states B → Bool ) → (fn : Naccept NFA finab nq x ≡ true )
|
|
203 → (qa : states A ) → ( (q : states A ∨ states B) → nq q ≡ true → ab-case q qa x )
|
|
204 → split (accept (automaton A) qa ) (contain B) x ≡ true
|
163
|
205 contain-A [] nq fn qa cond with found← finab fn
|
|
206 ... | S with found-q S | inspect found-q S | cond (found-q S) (bool-∧→tt-0 (found-p S))
|
|
207 ... | case1 qa' | record { eq = refl } | refl = bool-∧→tt-1 (found-p S)
|
|
208 ... | case2 qb | record { eq = refl } | ab = ⊥-elim ( ab (bool-∧→tt-1 (found-p S)))
|
|
209 contain-A (h ∷ t) nq fn qa cond with bool-≡-? ((aend (automaton A) qa) /\ accept (automaton B) (δ (automaton B) (astart B) h) t ) true
|
|
210 ... | yes eq = bool-or-41 eq
|
|
211 ... | no ne = bool-or-31 (contain-A t (Nmoves NFA finab nq h) fn (δ (automaton A) qa h) lemma11 ) where
|
|
212 lemma11 : (q : states A ∨ states B) → exists finab (λ qn → nq qn /\ Nδ NFA qn h q) ≡ true → ab-case q (δ (automaton A) qa h) t
|
|
213 lemma11 (case1 qa') ex with found← finab ex
|
|
214 ... | S with found-q S | inspect found-q S | cond (found-q S) (bool-∧→tt-0 (found-p S))
|
|
215 ... | case1 qa | record { eq = refl } | refl = sym ( equal→refl (afin A) ( bool-∧→tt-1 (found-p S) )) -- continued A case
|
|
216 ... | case2 qb | record { eq = refl } | nb with bool-∧→tt-1 (found-p S) -- δnfa (case2 q) i (case1 q₁) is false
|
|
217 ... | ()
|
|
218 lemma11 (case2 qb) ex with found← finab ex
|
|
219 ... | S with found-q S | inspect found-q S | cond (found-q S) (bool-∧→tt-0 (found-p S))
|
|
220 lemma11 (case2 qb) ex | S | case2 qb' | record { eq = refl } | nb = contra-position lemma13 nb where -- continued B case should fail
|
|
221 lemma13 : accept (automaton B) qb t ≡ true → accept (automaton B) qb' (h ∷ t) ≡ true
|
|
222 lemma13 fb = subst (λ k → accept (automaton B) k t ≡ true ) (sym (equal→refl (afin B) (bool-∧→tt-1 (found-p S)))) fb
|
|
223 lemma11 (case2 qb) ex | S | case1 qa | record { eq = refl } | refl with bool-∧→tt-1 (found-p S)
|
|
224 ... | eee = contra-position lemma12 ne where -- starting B case should fail
|
|
225 lemma12 : accept (automaton B) qb t ≡ true → aend (automaton A) qa /\ accept (automaton B) (δ (automaton B) (astart B) h) t ≡ true
|
|
226 lemma12 fb = bool-and-tt (bool-∧→tt-0 eee) (subst ( λ k → accept (automaton B) k t ≡ true ) (equal→refl (afin B) (bool-∧→tt-1 eee) ) fb )
|
141
|
227
|
163
|
228 lemma10 : Naccept NFA finab (equal? finab (case1 (astart A))) x ≡ true → split (contain A) (contain B) x ≡ true
|
|
229 lemma10 CC = contain-A x (Concat-NFA-start A B ) CC (astart A) lemma15 where
|
|
230 lemma15 : (q : states A ∨ states B) → Concat-NFA-start A B q ≡ true → ab-case q (astart A) x
|
|
231 lemma15 q nq=t with equal→refl finab nq=t
|
|
232 ... | refl = refl
|
141
|
233
|
163
|
234 closed-in-concat← : contain (M-Concat A B) x ≡ true → Concat (contain A) (contain B) x ≡ true
|
|
235 closed-in-concat← C with subset-construction-lemma← finab NFA (case1 (astart A)) x C
|
|
236 ... | CC = lemma10 CC
|
141
|
237
|
|
238
|
|
239
|
|
240
|