65
|
1 module regular-language where
|
|
2
|
|
3 open import Level renaming ( suc to Suc ; zero to Zero )
|
|
4 open import Data.List
|
|
5 open import Data.Nat hiding ( _≟_ )
|
70
|
6 open import Data.Fin hiding ( _+_ )
|
72
|
7 open import Data.Empty
|
65
|
8 open import Data.Product
|
|
9 -- open import Data.Maybe
|
|
10 open import Relation.Nullary
|
|
11 open import Relation.Binary.PropositionalEquality hiding ( [_] )
|
|
12 open import logic
|
70
|
13 open import nat
|
65
|
14 open import automaton
|
|
15 open import finiteSet
|
|
16
|
|
17 language : { Σ : Set } → Set
|
|
18 language {Σ} = List Σ → Bool
|
|
19
|
|
20 language-L : { Σ : Set } → Set
|
|
21 language-L {Σ} = List (List Σ)
|
|
22
|
|
23 open Automaton
|
|
24
|
|
25 record RegularLanguage ( Σ : Set ) : Set (Suc Zero) where
|
|
26 field
|
|
27 states : Set
|
|
28 astart : states
|
70
|
29 aℕ : ℕ
|
|
30 afin : FiniteSet states {aℕ}
|
65
|
31 automaton : Automaton states Σ
|
|
32 contain : List Σ → Bool
|
|
33 contain x = accept automaton astart x
|
|
34
|
|
35 Union : {Σ : Set} → ( A B : language {Σ} ) → language {Σ}
|
|
36 Union {Σ} A B x = (A x ) \/ (B x)
|
|
37
|
|
38 split : {Σ : Set} → (List Σ → Bool)
|
|
39 → ( List Σ → Bool) → List Σ → Bool
|
|
40 split x y [] = x [] /\ y []
|
|
41 split x y (h ∷ t) = (x [] /\ y (h ∷ t)) \/
|
|
42 split (λ t1 → x ( h ∷ t1 )) (λ t2 → y t2 ) t
|
|
43
|
|
44 Concat : {Σ : Set} → ( A B : language {Σ} ) → language {Σ}
|
|
45 Concat {Σ} A B = split A B
|
|
46
|
|
47 {-# TERMINATING #-}
|
|
48 Star : {Σ : Set} → ( A : language {Σ} ) → language {Σ}
|
|
49 Star {Σ} A = split A ( Star {Σ} A )
|
|
50
|
69
|
51 test-split : {Σ : Set} → {A B : List In2 → Bool} → split A B ( i0 ∷ i1 ∷ i0 ∷ [] ) ≡ (
|
|
52 ( A [] /\ B ( i0 ∷ i1 ∷ i0 ∷ [] ) ) \/
|
|
53 ( A ( i0 ∷ [] ) /\ B ( i1 ∷ i0 ∷ [] ) ) \/
|
|
54 ( A ( i0 ∷ i1 ∷ [] ) /\ B ( i0 ∷ [] ) ) \/
|
|
55 ( A ( i0 ∷ i1 ∷ i0 ∷ [] ) /\ B [] )
|
|
56 )
|
|
57 test-split {_} {A} {B} = refl
|
|
58
|
65
|
59 open RegularLanguage
|
|
60 isRegular : {Σ : Set} → (A : language {Σ} ) → ( x : List Σ ) → (r : RegularLanguage Σ ) → Set
|
|
61 isRegular A x r = A x ≡ contain r x
|
|
62
|
73
|
63 postulate
|
|
64 fin-× : {A B : Set} → { a b : ℕ } → FiniteSet A {a} → FiniteSet B {b} → FiniteSet (A × B) {a * b}
|
|
65
|
65
|
66 M-Union : {Σ : Set} → (A B : RegularLanguage Σ ) → RegularLanguage Σ
|
|
67 M-Union {Σ} A B = record {
|
|
68 states = states A × states B
|
|
69 ; astart = ( astart A , astart B )
|
73
|
70 ; aℕ = aℕ A * aℕ B
|
|
71 ; afin = fin-× (afin A) (afin B)
|
65
|
72 ; automaton = record {
|
|
73 δ = λ q x → ( δ (automaton A) (proj₁ q) x , δ (automaton B) (proj₂ q) x )
|
|
74 ; aend = λ q → ( aend (automaton A) (proj₁ q) \/ aend (automaton B) (proj₂ q) )
|
|
75 }
|
|
76 }
|
|
77
|
|
78 closed-in-union : {Σ : Set} → (A B : RegularLanguage Σ ) → ( x : List Σ ) → isRegular (Union (contain A) (contain B)) x ( M-Union A B )
|
|
79 closed-in-union A B [] = lemma where
|
|
80 lemma : aend (automaton A) (astart A) \/ aend (automaton B) (astart B) ≡
|
|
81 aend (automaton A) (astart A) \/ aend (automaton B) (astart B)
|
|
82 lemma = refl
|
|
83 closed-in-union {Σ} A B ( h ∷ t ) = lemma1 t ((δ (automaton A) (astart A) h)) ((δ (automaton B) (astart B) h)) where
|
|
84 lemma1 : (t : List Σ) → (qa : states A ) → (qb : states B ) →
|
|
85 accept (automaton A) qa t \/ accept (automaton B) qb t
|
|
86 ≡ accept (automaton (M-Union A B)) (qa , qb) t
|
|
87 lemma1 [] qa qb = refl
|
|
88 lemma1 (h ∷ t ) qa qb = lemma1 t ((δ (automaton A) qa h)) ((δ (automaton B) qb h))
|
|
89
|
|
90 -- M-Concat : {Σ : Set} → (A B : RegularLanguage Σ ) → RegularLanguage Σ
|
|
91 -- M-Concat {Σ} A B = record {
|
|
92 -- states = states A ∨ states B
|
|
93 -- ; astart = case1 (astart A )
|
|
94 -- ; automaton = record {
|
|
95 -- δ = {!!}
|
|
96 -- ; aend = {!!}
|
|
97 -- }
|
|
98 -- }
|
|
99 --
|
|
100 -- closed-in-concat : {Σ : Set} → (A B : RegularLanguage Σ ) → ( x : List Σ ) → isRegular (Concat (contain A) (contain B)) x ( M-Concat A B )
|
|
101 -- closed-in-concat = {!!}
|
70
|
102
|
|
103 open import nfa
|
|
104 open import sbconst2
|
|
105 open FiniteSet
|
86
|
106 open import Data.Nat.Properties hiding ( _≟_ )
|
70
|
107 open import Relation.Binary as B hiding (Decidable)
|
|
108
|
73
|
109 postulate
|
|
110 fin-∨ : {A B : Set} → { a b : ℕ } → FiniteSet A {a} → FiniteSet B {b} → FiniteSet (A ∨ B) {a + b}
|
|
111 fin→ : {A : Set} → { a : ℕ } → FiniteSet A {a} → FiniteSet (A → Bool ) {exp 2 a}
|
70
|
112
|
|
113 Concat-NFA : {Σ : Set} → (A B : RegularLanguage Σ ) → NAutomaton (states A ∨ states B) Σ
|
86
|
114 Concat-NFA {Σ} A B = record { Nδ = δnfa ; Nend = nend }
|
|
115 module Concat-NFA where
|
70
|
116 δnfa : states A ∨ states B → Σ → states A ∨ states B → Bool
|
|
117 δnfa (case1 q) i (case1 q₁) = equal? (afin A) (δ (automaton A) q i) q₁
|
|
118 δnfa (case1 qa) i (case2 qb) = (aend (automaton A) qa ) /\ (equal? (afin B) qb (astart B) )
|
|
119 δnfa (case2 q) i (case2 q₁) = equal? (afin B) (δ (automaton B) q i) q₁
|
|
120 δnfa _ i _ = false
|
|
121 nend : states A ∨ states B → Bool
|
|
122 nend (case2 q) = aend (automaton B) q
|
|
123 nend _ = false
|
|
124
|
75
|
125 -- Concat-NFA-start : {Σ : Set} → (A B : RegularLanguage Σ ) → states A ∨ states B → Bool
|
|
126 -- Concat-NFA-start A B (case1 q) = equal? (afin A) q (astart A)
|
|
127 -- Concat-NFA-start _ _ _ = false
|
|
128
|
70
|
129 Concat-NFA-start : {Σ : Set} → (A B : RegularLanguage Σ ) → states A ∨ states B → Bool
|
75
|
130 Concat-NFA-start A B q = equal? (fin-∨ (afin A) (afin B)) (case1 (astart A)) q
|
70
|
131
|
|
132 M-Concat : {Σ : Set} → (A B : RegularLanguage Σ ) → RegularLanguage Σ
|
|
133 M-Concat {Σ} A B = record {
|
|
134 states = states A ∨ states B → Bool
|
|
135 ; astart = Concat-NFA-start A B
|
|
136 ; aℕ = finℕ finf
|
|
137 ; afin = finf
|
|
138 ; automaton = subset-construction fin (Concat-NFA A B) (case1 (astart A))
|
|
139 } where
|
|
140 fin : FiniteSet (states A ∨ states B ) {aℕ A + aℕ B}
|
|
141 fin = fin-∨ (afin A) (afin B)
|
|
142 finf : FiniteSet (states A ∨ states B → Bool )
|
|
143 finf = fin→ fin
|
|
144
|
72
|
145 record Split {Σ : Set} (A : List Σ → Bool ) ( B : List Σ → Bool ) (x : List Σ ) : Set where
|
|
146 field
|
|
147 sp0 : List Σ
|
|
148 sp1 : List Σ
|
|
149 sp-concat : sp0 ++ sp1 ≡ x
|
|
150 prop0 : A sp0 ≡ true
|
|
151 prop1 : B sp1 ≡ true
|
|
152
|
|
153 open Split
|
|
154
|
|
155 list-empty++ : {Σ : Set} → (x y : List Σ) → x ++ y ≡ [] → (x ≡ [] ) ∧ (y ≡ [] )
|
|
156 list-empty++ [] [] refl = record { proj1 = refl ; proj2 = refl }
|
|
157 list-empty++ [] (x ∷ y) ()
|
|
158 list-empty++ (x ∷ x₁) y ()
|
|
159
|
|
160 open _∧_
|
|
161
|
74
|
162 open import Relation.Binary.PropositionalEquality hiding ( [_] )
|
71
|
163
|
74
|
164 c-split-lemma : {Σ : Set} → (A B : List Σ → Bool ) → (h : Σ) → ( t : List Σ ) → split A B (h ∷ t ) ≡ true
|
|
165 → ( ¬ (A [] ≡ true )) ∨ ( ¬ (B ( h ∷ t ) ≡ true) )
|
|
166 → split (λ t1 → A (h ∷ t1)) B t ≡ true
|
|
167 c-split-lemma {Σ} A B h t eq (case1 ¬p ) = sym ( begin
|
|
168 true
|
|
169 ≡⟨ sym eq ⟩
|
|
170 split A B (h ∷ t )
|
|
171 ≡⟨⟩
|
|
172 A [] /\ B (h ∷ t) \/ split (λ t1 → A (h ∷ t1)) B t
|
|
173 ≡⟨ cong ( λ k → k \/ split (λ t1 → A (h ∷ t1)) B t ) (bool-and-1 (¬-bool-t ¬p)) ⟩
|
|
174 false \/ split (λ t1 → A (h ∷ t1)) B t
|
|
175 ≡⟨ bool-or-1 refl ⟩
|
|
176 split (λ t1 → A (h ∷ t1)) B t
|
|
177 ∎ ) where open ≡-Reasoning
|
|
178 c-split-lemma {Σ} A B h t eq (case2 ¬p ) = sym ( begin
|
|
179 true
|
|
180 ≡⟨ sym eq ⟩
|
|
181 split A B (h ∷ t )
|
|
182 ≡⟨⟩
|
|
183 A [] /\ B (h ∷ t) \/ split (λ t1 → A (h ∷ t1)) B t
|
|
184 ≡⟨ cong ( λ k → k \/ split (λ t1 → A (h ∷ t1)) B t ) (bool-and-2 (¬-bool-t ¬p)) ⟩
|
|
185 false \/ split (λ t1 → A (h ∷ t1)) B t
|
|
186 ≡⟨ bool-or-1 refl ⟩
|
|
187 split (λ t1 → A (h ∷ t1)) B t
|
|
188 ∎ ) where open ≡-Reasoning
|
|
189
|
|
190 c-split : {Σ : Set} → (A B : List Σ → Bool ) → ( x : List Σ ) → split A B x ≡ true → Split A B x
|
|
191 c-split {Σ} A B [] eq with bool-≡-? (A []) true | bool-≡-? (B []) true
|
73
|
192 c-split {Σ} A B [] eq | yes eqa | yes eqb =
|
|
193 record { sp0 = [] ; sp1 = [] ; sp-concat = refl ; prop0 = eqa ; prop1 = eqb }
|
|
194 c-split {Σ} A B [] eq | yes p | no ¬p = ⊥-elim (lemma-∧-1 eq (¬-bool-t ¬p ))
|
|
195 c-split {Σ} A B [] eq | no ¬p | t = ⊥-elim (lemma-∧-0 eq (¬-bool-t ¬p ))
|
74
|
196 c-split {Σ} A B (h ∷ t ) eq with bool-≡-? (A []) true | bool-≡-? (B (h ∷ t )) true
|
|
197 ... | yes px | yes py = record { sp0 = [] ; sp1 = h ∷ t ; sp-concat = refl ; prop0 = px ; prop1 = py }
|
|
198 ... | no px | _ with c-split (λ t1 → A ( h ∷ t1 )) B t (c-split-lemma A B h t eq (case1 px) )
|
|
199 ... | S = record { sp0 = h ∷ sp0 S ; sp1 = sp1 S ; sp-concat = cong ( λ k → h ∷ k) (sp-concat S) ; prop0 = prop0 S ; prop1 = prop1 S }
|
|
200 c-split {Σ} A B (h ∷ t ) eq | _ | no px with c-split (λ t1 → A ( h ∷ t1 )) B t (c-split-lemma A B h t eq (case2 px) )
|
|
201 ... | S = record { sp0 = h ∷ sp0 S ; sp1 = sp1 S ; sp-concat = cong ( λ k → h ∷ k) (sp-concat S) ; prop0 = prop0 S ; prop1 = prop1 S }
|
|
202
|
|
203 split++ : {Σ : Set} → (A B : List Σ → Bool ) → ( x y : List Σ ) → A x ≡ true → B y ≡ true → split A B (x ++ y ) ≡ true
|
|
204 split++ {Σ} A B [] [] eqa eqb = begin
|
|
205 split A B []
|
|
206 ≡⟨⟩
|
|
207 A [] /\ B []
|
|
208 ≡⟨ cong₂ (λ j k → j /\ k ) eqa eqb ⟩
|
|
209 true
|
|
210 ∎ where open ≡-Reasoning
|
|
211 split++ {Σ} A B [] (h ∷ y ) eqa eqb = begin
|
|
212 split A B (h ∷ y )
|
|
213 ≡⟨⟩
|
|
214 A [] /\ B (h ∷ y) \/ split (λ t1 → A (h ∷ t1)) B y
|
|
215 ≡⟨ cong₂ (λ j k → j /\ k \/ split (λ t1 → A (h ∷ t1)) B y ) eqa eqb ⟩
|
|
216 true /\ true \/ split (λ t1 → A (h ∷ t1)) B y
|
|
217 ≡⟨⟩
|
|
218 true \/ split (λ t1 → A (h ∷ t1)) B y
|
|
219 ≡⟨⟩
|
|
220 true
|
|
221 ∎ where open ≡-Reasoning
|
|
222 split++ {Σ} A B (h ∷ t) y eqa eqb = begin
|
|
223 split A B ((h ∷ t) ++ y)
|
|
224 ≡⟨⟩
|
|
225 A [] /\ B (h ∷ t ++ y) \/ split (λ t1 → A (h ∷ t1)) B (t ++ y)
|
|
226 ≡⟨ cong ( λ k → A [] /\ B (h ∷ t ++ y) \/ k ) ( begin
|
|
227 split (λ t1 → A (h ∷ t1)) B (t ++ y)
|
|
228 ≡⟨ split++ {Σ} (λ t1 → A (h ∷ t1)) B t y eqa eqb ⟩
|
|
229 true
|
|
230 ∎ ) ⟩
|
|
231 A [] /\ B (h ∷ t ++ y) \/ true
|
|
232 ≡⟨ bool-or-3 ⟩
|
|
233 true
|
|
234 ∎ where open ≡-Reasoning
|
|
235
|
86
|
236 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality n n -- (Level.suc n) already in finiteSet
|
75
|
237
|
|
238 open NAutomaton
|
70
|
239
|
71
|
240 closed-in-concat : {Σ : Set} → (A B : RegularLanguage Σ ) → ( x : List Σ ) → isRegular (Concat (contain A) (contain B)) x ( M-Concat A B )
|
75
|
241 closed-in-concat {Σ} A B x = ≡-Bool-func lemma3 lemma4 where
|
76
|
242 finav = (fin-∨ (afin A) (afin B))
|
|
243 NFA = (Concat-NFA A B)
|
|
244 abmove : (q : states A ∨ states B) → (h : Σ ) → states A ∨ states B
|
|
245 abmove (case1 q) h = case1 (δ (automaton A) q h)
|
|
246 abmove (case2 q) h = case2 (δ (automaton B) q h)
|
|
247 nmove : (q : states A ∨ states B) (nq : states A ∨ states B → Bool ) → (nq q ≡ true) → ( h : Σ ) →
|
|
248 exists finav (λ qn → nq qn /\ Nδ NFA qn h (abmove q h)) ≡ true
|
86
|
249 nmove (case1 q) nq nqt h = found finav {_} {(case1 q)} ( bool-and-tt nqt lemma-nmove-a ) where
|
|
250 lemma-nmove-a : Nδ NFA (case1 q) h (abmove (case1 q) h) ≡ true
|
|
251 lemma-nmove-a with F←Q (afin A) (δ (automaton A) q h) ≟ F←Q (afin A) (δ (automaton A) q h)
|
|
252 lemma-nmove-a | yes refl = refl
|
|
253 lemma-nmove-a | no ne = ⊥-elim (ne refl)
|
|
254 nmove (case2 q) nq nqt h = found finav {_} {(case2 q)} ( bool-and-tt nqt lemma-nmove ) where
|
|
255 lemma-nmove : Nδ NFA (case2 q) h (abmove (case2 q) h) ≡ true
|
|
256 lemma-nmove with F←Q (afin B) (δ (automaton B) q h) ≟ F←Q (afin B) (δ (automaton B) q h)
|
|
257 lemma-nmove | yes refl = refl
|
|
258 lemma-nmove | no ne = ⊥-elim (ne refl)
|
76
|
259 lemma6 : (z : List Σ) → (q : states B) → (nq : states A ∨ states B → Bool ) → (nq (case2 q) ≡ true) → ( accept (automaton B) q z ≡ true )
|
|
260 → Naccept NFA finav nq z ≡ true
|
86
|
261 lemma6 [] q nq nqt fb = lemma8 where
|
76
|
262 lemma8 : exists finav ( λ q → nq q /\ Nend NFA q ) ≡ true
|
86
|
263 lemma8 = found finav {_} {case2 q} ( bool-and-tt nqt fb )
|
76
|
264 lemma6 (h ∷ t ) q nq nq=q fb = lemma6 t (δ (automaton B) q h) (Nmoves NFA finav nq h) (nmove (case2 q) nq nq=q h) fb
|
|
265 lemma7 : (y z : List Σ) → (q : states A) → (nq : states A ∨ states B → Bool ) → (nq (case1 q) ≡ true)
|
75
|
266 → ( accept (automaton A) q y ≡ true ) → ( accept (automaton B) (astart B) z ≡ true )
|
76
|
267 → Naccept NFA finav nq (y ++ z) ≡ true
|
86
|
268 lemma7 [] z q nq nq=q fa fb = lemma6 z (astart B) nq lemma71 fb where
|
|
269 lemma71 : nq (case2 (astart B)) ≡ true
|
|
270 lemma71 = {!!}
|
|
271 lemma-nq=q : (nq (case1 q) ≡ true)
|
|
272 lemma-nq=q = nq=q
|
76
|
273 lemma7 (h ∷ t) z q nq nq=q fa fb = lemma7 t z (δ (automaton A) q h) (Nmoves NFA finav nq h) (nmove (case1 q) nq nq=q h) fa fb where
|
|
274 lemma9 : equal? finav (case1 (astart A)) (case1 (astart A)) ≡ true
|
|
275 lemma9 with Data.Fin._≟_ (F←Q finav (case1 (astart A))) ( F←Q finav (case1 (astart A)) )
|
|
276 lemma9 | yes refl = refl
|
|
277 lemma9 | no ¬p = ⊥-elim ( ¬p refl )
|
75
|
278 lemma5 : Split (contain A) (contain B) x
|
76
|
279 → Naccept NFA finav (equal? finav (case1 (astart A))) x ≡ true
|
|
280 lemma5 S = subst ( λ k → Naccept NFA finav (equal? finav (case1 (astart A))) k ≡ true ) ( sp-concat S )
|
|
281 (lemma7 (sp0 S) (sp1 S) (astart A) (equal? finav (case1 (astart A))) lemma9 (prop0 S) (prop1 S) )
|
71
|
282 lemma3 : Concat (contain A) (contain B) x ≡ true → contain (M-Concat A B) x ≡ true
|
74
|
283 lemma3 concat with c-split (contain A) (contain B) x concat
|
75
|
284 ... | S = begin
|
76
|
285 accept (subset-construction finav NFA (case1 (astart A))) (Concat-NFA-start A B ) x
|
|
286 ≡⟨ ≡-Bool-func (subset-construction-lemma← finav NFA (case1 (astart A)) x )
|
|
287 (subset-construction-lemma→ finav NFA (case1 (astart A)) x ) ⟩
|
|
288 Naccept NFA finav (equal? finav (case1 (astart A))) x
|
75
|
289 ≡⟨ lemma5 S ⟩
|
|
290 true
|
|
291 ∎ where open ≡-Reasoning
|
71
|
292 lemma4 : contain (M-Concat A B) x ≡ true → Concat (contain A) (contain B) x ≡ true
|
74
|
293 lemma4 C = {!!} -- split++ (contain A) (contain B) x y (accept ?) (accept ?)
|
71
|
294
|
|
295
|