Unified Decision Procedures for Regular Expression Equivalence

Size: px
Start display at page:

Download "Unified Decision Procedures for Regular Expression Equivalence"

Transcription

1 Unified Decision Procedures for Regulr Expression Equivlence Tois Nipkow nd Dmitriy Trytel Fkultät für Informtik, Technische Universität München, Germny Astrct. We formlize unified frmework for verified decision procedures for regulr expression equivlence. Five recently pulished formliztions of such decision procedures (three sed on derivtives, two on mrked regulr expressions) cn e otined s instnces of the frmework. We discover tht the two pproches sed on mrked regulr expressions, which were previously thought to e the sme, re different, nd we prove quotient reltion etween the utomt produced y them. The common frmework mkes it possile to compre the performnce of the different decision procedures in meningful wy. 1 Introduction Equivlence of regulr expressions is perennil topic in computer science. Recently it hs spwned numer of formlized nd verified decision procedures for this tsk in interctive theorem provers [3, 6, 10, 19, 21]. Except for the formliztion y Brint nd Pous [6], ll these decision procedures operte directly on vritions of regulr expressions. Although they (implicitly) uild utomt, the sttes of the utomt re leled with regulr expressions, nd there is no glol trnsition tle ut the nextstte function is computle from the regulr expressions. The motivtion for working with regulr expressions is simplicity: regulr expressions re free dttype which proof ssistnts nd their users love ecuse it mens induction, recursion nd equtionl resoning the core competence of proof ssistnts nd functionl progrmming lnguges. Yet ll these decision procedures sed on regulr expressions look very different. Of course, the next-stte functions ll differ, ut so do the ctul decision procedures nd their correctness, completeness nd termintion proofs. The contriutions of our pper re the following: A unified frmework (Sect. 3) tht we instntite with ll the ove pproches (Sects. 4 nd 5). The frmework is simple reflexive trnsitive closure computtion tht enumertes the sttes of product utomton. Proofs of correctness, completeness nd termintion tht re performed once nd for ll for the frmework sed on few properties of the next-stte function. A new perspective on prtil derivtives tht recsts them s Brzozowski derivtives followed y some rewriting (Sect. 4). The discovery tht Asperti s lgorithm is not the one y McNughton-Ymd [20], s stted y Asperti [3], ut dul construction which pprently hd not een considered in the literture nd which produces smller utomt (Sect. 5). An empiricl comprison of the performnce of the different pproches (Sect. 6). The discussion of relted work is distriuted over the relevnt sections of the pper.

2 2 Preliminries Iselle/HOL is sed on Church s simple type theory (see [22, Prt I] for recent introduction). Types τ re uilt from type vriles α, β, etc. vi function types, other type constructors re written postfix. The nottion t :: τ mens tht term t hs type τ. Types αset nd αlist re the types of sets nd lists of elements of type α. They come with the following voculry: function set (conversion from lists to sets), [] (empty list), # (list (ppend), hd (hed), tl (til) nd mp. Recursive functions over dttypes re executle, nd Iselle cn generte from them code in functionl lnguges [15]. This includes functions on finite sets [14]. Unless stted otherwise ll functions in this pper re executle. Locles [4] re Iselle s tool for modelling prmeterized systems. A locle fixes prmeters nd sttes ssumptions out them: locle A = fixes x 1 nd... nd x n ssumes n 1 : P 1 x nd... nd n m : P m x In the context of the locle A, we cn define constnts tht depend on the prmeters x i nd prove properties out those constnts using the ssumptions P i (ccessed under the nme n i ). Prmeters cn e instntited: interprettion J: A where x 1 = t 1... x n = t n. The commnd issues proof oligtions P i t (tht the user must dischrge) nd exports constnts nd theorems from the locle with x i instntited to t i. Multiple interprettions of the sme locle re possile; the prefix J. dismigutes different instnces. Regulr expressions re defined s recursive dttype: dttype αrexp = 0 1 A α αrexp + αrexp αrexp αrexp (αrexp) with the usul (non-executle) semntics L :: α rexp α lng, where α lng is short for (α list) set. In concrete regulr expressions, we sometimes omit the constructor A for redility. The recursive function nullle :: α rexp ool stisfies nullle r [] L r. The functions Σ :: αrexp αset nd toms :: αrexp αlist compute the set nd list of toms (the rguments of constructor A) in regulr expression. The (non-executle) left quotient of lnguge L :: α lng w.r.t. some :: α is defined y D L = {w # w L}. The extension of D from single symols to words w :: αlist cn e expressed s fold D w where fold :: (α β β) αlist β β. 3 Regulr Expression Equivlence Frmework Regulr expression (lnguge) equivlence is usully reduced to (lnguge) equivlence of utomt. In principle our frmework does the sme, except tht we construct the utomt on the fly nd replce the trditionl trnsition tle y computtions on regulr-expression-like ojects. We strt y relting regulr expressions nd utomt. Left quotients of regulr lnguge L cn e understood s sttes of deterministic utomton M L with the initil stte L, the trnsition function D, nd the ccepting sttes eing those lnguges K for which [] K holds. This utomton (restricted to rechle sttes) is finite nd miniml y the Myhill-Nerode theorem. 1 The following locle cptures this left-quotient-sed view of n utomton: 1 Note tht the Myhill-Nerode reltion L cn e defined s v L w fold D v L = fold D w L. The quotient of M L y this reltion is isomorphic to M L ; hence M L is miniml.

3 locle rexpda = fixes ι :: αrexp σ nd L :: σ αlng nd δ :: α σ σ nd o :: σ ool ssumes ιl: L(ι r) = L r nd δl: L(δ s) = D (L s) nd ol: o s [] L s The prmeters ι nd L formlize wht regulr-expression-like mens: ι r emeds the regulr expression r into stte of type σ, wheres L gives elements of σ lnguge semntics, which coincides with the lnguge semntics of regulr expressions y the ssumption ιl. The function δ is the symolic computtion of left quotients on σ ccording to δl. It cn e regrded s the trnsition function of n utomton with sttes in σ nd the initil stte ι r. Accepting sttes of this utomton re given y o. Let us develop nd verify some lgorithms in the context of rexpda. For strt, regulr expression mtching is esy to define mtch r w = o (fold δ w (ι r)) nd prove correct: mtch r w w L r. Now we tckle the equivlence checker. We follow the well-known product utomton construction where lnguge equivlence mens o s 1 o s 2 for ll sttes (s 1, s 2 ) of the product utomton. Alterntively, one cn view this procedure s the construction of isimultion reltion etween two utomt: lnguge equivlence nd existence of isimultion coincide for deterministic utomt [24]. The set of rechle sttes of n utomton cn e otined s the reflexive trnsitive closure of the strt stte under λ p. mp (λ. δ p) s where s :: αlist is the lphet. We define reflexive trnsitive closure opertion rtc :: (α ool) (α αlist) α (αlist αset)option where type αoption is the dttype None Some α. It is used to encode whether the closure is finite (Some is returned) or infinite (None is returned). The function rtc is defined using while comintor nd is executle (provided its rguments eing executle); the result Some corresponds to terminting computtion [19]. The definition cn e found in Iselle/HOL s lirry theory While_ Comintor under its full nme rtrncl_ while. The prmeters nd result of rtc p next strt hve the following mening: Predicte p is test tht stops the closure computtion if n element not stisfying p is found; this is merely n optimiztion. Function next mps n element to list of successors. Of course strt is the strt element. A result Some (ws, Z) mens tht the closure computtion terminted with worklist ws nd set of rechle elements Z. If ws is empty, Z is the set of ll elements rechle from strt; otherwise, the computtion ws stopped ecuse n element not stisfying p ws found. More precisely, we proved rtc p next strt = Some (ws, Z) = if ws = [] then Z = R ( z Z. p z) else p (hd ws) hd ws R (1) where R = {(x, y) y set (next x)} {strt} nd is infix reltion ppliction: r {x} = {y (x, y) r}. The stte spce of the product utomton is computed s follows: closure :: αlist σ σ ((σ σ)list (σ σ)set)option closure s = rtc (λ(s, t). o s ot) (λ(s, t). mp (λ. (δ s, δ t)) s)

4 The predicte λ(s, t). o s ot stops the computtion s soon s contrdiction to lnguge equlity is found. The ctul lnguge equivlence checker merely needs to test if the worklist is empty t the end: eqv :: αrexp αrexp ool eqv r s = cse closure (toms toms s) (ι r, ι s) of Some ([], _) True _ Flse The lphet given to closure is the conctention of the toms in the two expressions. Soundness of eqv is n esy consequence of the following property, which in turn follows from (1): closure (toms toms s) (ι r,ι s) = Some (ws, Z) = ws = [] L r = L s Theorem 1 (in rexpda). eqv r s = L r = L s. This is prtil correctness sttement ecuse it ssumes tht the cll to closure in eqv returns Some, i.e. termintes. Termintion of closure needs finiteness of the underlying utomton. Therefore we extend rexpda with n explicit ssumption of finiteness: locle rexpdfa = rexpda + ssumes fin: finite {fold δ w (ι r) w :: αlist} In this context the termintion lemm for closure is n esy consequence of fin nd the following termintion property of rtc: finite ({(x, y) y set ( f x)} {x}) = y. rtc p f x = Some y Lemm 2 (in rexpdfa). closure s (ι r, ι s) None. Together with (2) this implies completeness of eqv: Theorem 3 (in rexpdfa). L r = L s = eqv r s. This is the end of ll considertions out equivlence of regulr expressions. The rest of the pper merely needs to focus on vrious methods for turning regulr expressions into finite utomt in the sense of rexpdfa. Note tht M L defined ove constitutes first vlid interprettion of rexpdfa. The proof of fin requires the Myhill-Nerode theorem. interprettion M: rexpdfa where ι r = L r δ L = D L o L = [] L L L = L This interprettion is not executle ecuse neither its next-step function D (eing sed on infinite sets of words defined y set comprehension) nor the equlity on σ = αlng (which is needed for the closure computtion) is executle. (2)

5 4 Derivtives In 1964, Brzozowski [7] showed how to compute left quotients syntcticlly s derivtives of regulr expressions. Derivtives hve een rediscovered in proof ssistnts y Kruss nd Nipkow [19] nd Coqund nd Siles [10]. Our first executle instntitions of the frmework reuse infrstructure from erlier formliztions in Iselle [19, 26]. A refinement of Brzozowski s pproch, prtil derivtives, ws introduced y Antimirov [2] nd formlized y Moreir et l. [21] in Coq nd y Wu et l. [27] in Iselle. Prtil derivtives operte on finite sets of regulr expressions. They cn e viewed either s nondeterministic utomton with regulr expressions s sttes or s the corresponding deterministic utomton otined y the suset construction. In the following, we integrte the two notions in our frmework nd show how derivtives cn e used to simulte prtil derivtives without invoking sets explicitly. 4.1 Brzozowski s Derivtives Given letter c nd regulr expression r, the (Brzozowski) derivtive der :: α αrexp αrexp of r w.r.t. is defined y primitive recursion: der _ 0 = 0 der _ 1 = 0 der (A x) = if x = then 1 else 0 der (r + s) = der r + der s der (r s) = if nullle r then (der r s) + der s else der r s der (r ) = der r r It follows y induction on r tht the lnguge of the derivtive der r is exctly the left quotient D (L r). This property corresponds exctly to the ssumption δl of the locle rexpda. Hence it suggests the following interprettion: interprettion rexpda where ι r = r δ r = der r o r = nullle r L r = L r Unfortuntely, the sound equivlence checker tht is produced y this interprettion is useless in prctice, ecuse it will rrely terminte. For exmple, the utomton constructed from the regulr expression is infinite, s ll derivtives w.r.t. words n re distinct: fold der 1 = 1 ; fold der n+1 = 0 + fold der n. Fortuntely, Brzozowski showed tht there re finitely mny equivlence clsses of derivtives modulo ssocitivity, commuttivity nd idempotence (ACI) of the + constructor. We prove tht the numer of distinct derivtives of r modulo ACI is finite: finite {[fold der w r] w (Σ r) } where [r] = {s r s} denotes the equivlence clss of r nd the ACI equivlence is defined inductively s follows. r + (s + t) (r + s) + t r + s s + r r + r r r s r s s t r r s r r t r 1 s 1 r 2 s 2 r 1 s 1 r 2 s 2 r s r 1 + r 2 s 1 + s 2 r 1 r 2 s 1 s 2 r s ACI-equivlent regulr expressions r s hve the sme toms nd sme lnguges, nd their equivlence is preserved y the derivtive: der r der s for ll Σr.

6 This enles the following interprettion tht opertes on ACI equivlence clsses. We otin first totlly correct nd complete equivlence checker D.eqv in Iselle/HOL. interprettion D : rexpdfa where ι r = [r] δ [r] = [der r] o[r] = nullle r [(1 ) + 0] L[r] = L r [((0 + 1 ) + 0) + 0] [((0 + 0 ) + 1) + 0] [ ] [(0 ) + 1] [((0 + 0 ) + 0) + 0] Fig. 1: Derivtive utomton modulo ACI for Techniclly, the formliztion defines quotient type [18] of regulr expressions modulo ACI to represent equivlence clsses nd uses Lifting nd Trnsfer [17] to lift opertions on regulr expressions to opertions on equivlence clsses. The ove presenttion of definitions of the locle prmeters y pttern mtching on equivlence clsses resemles the code generted y Iselle for quotients ( pseudo-constructor [14], [_], wrps concrete representtive r), rther thn the ctul definitions y Lifting. Since the equivlence checker must compre equivlence clsses, the code genertion for quotients requires n executle equlity (i.e. decision procedure for - equivlence). We chieve this through n ACI normliztion function _ tht mps regulr expression r to cnonicl representtive of [r] y sorting ll summnds w.r.t. n ritrry fixed liner order while removing duplictes. The definition of _ employs smrt (simplifying) constructor, whose equtions re mtched sequentilly. 0 = 0 1 = 1 A = A r + s = r s r s = r s r = r (r + s) t = r (s t) r (s + t) = if r = s then s + t else if r s then r + (s + t) else s + (r t) r s = if r = s then r else if r s then r + s else s + r We otin n executle decision procedure for ACI equivlence: r s r = s. This mkes D.eqv executle, yielding verified code in different functionl progrmming lnguges vi Iselle s code genertor. Yet, the performnce of the generted code is disppointing. Fig. 1 shows why: Derivtions clutter concrete representtives with duplicted summnds. Further derivtion steps perform the sme computtion repetedly nd hence ecome incresingly expensive. This ottleneck is voided y tking cnonicl ACI-normlized representtives s sttes yielding second interprettion. interprettion D: rexpdfa where ι r = r δ r = der r o r = nullle r 0 + (1 ) L r = L r 1 + (0 ) 0 + (0 + 1 ) 0 + (1 + (0 ) ) 0 + (0 ) Fig. 2: ACI-normlized derivtive utomton for

7 A few points re worth mentioning here: First, D does not use the quotient type it opertes directly on cnonicl representtives nd therefore cn use structurl equlity for comprison (rther thn ). Second, the interprettions D nd D yield structurlly the sme utomt, lthough with different lels. Fig. 2 shows the utomton produced y D for. This oservtion which enles us to reuse the techniclly involved proof of D.fin to dischrge D.fin relies crucilly on our normliztion function _ eing idempotent nd well-ehved w.r.t. derivtives: Lemm 4. We hve r = r nd der r = der r for ll Σ r. The utomton from Fig. 2 shows tht the stte lels still contin superfluous informtion, notly in the form of 0s nd 1s. A corser reltion thn -equivlence, denoted, dresses this concern. We omit the strightforwrd inductive definition of, which cncels 0s nd 1s where possile nd tkes the ssocitivity of conctention into ccount. Corseness ([r] [r] ) together with D.fin implies finiteness of equivlence clsses of derivtives modulo : finite {[fold der w r] w (Σ r) }. As efore, to void working with equivlence clsses, we use recursively defined -normliztion function _ similr to _ (it corresponds to the norm function from the formliztion y Kruss nd Nipkow [19]). However, _ (lso ) is not well-ehved w.r.t. derivtives: for exmple, der (( + 1) ( )) der ((( + 1) ( )) ). The normliztion would need to tke the distriutivity of over + into ccount to prevent this disequlity, ut even with this ddition forml proof of well-ehvedness seems difficult. Furthermore, our evlution (Sect. 6) suggests tht not too much energy should e invested in finding this proof. Thus, the following interprettion gives only prtil correctness result. interprettion N: rexpda where ι r = r δ r = der r o r = nullle r L r = L r 1 0 Fig. 3: Normlized derivtive utomton for In prctice, we did not find n input for which N would construct n infinite utomton. For the exmple it even yields the miniml utomton shown in Fig Prtil Derivtives Prtil derivtives split certin +-constructors into sets of regulr expressions, thus cpturing ACI equivlence directly in the dt structure. The utomton construction for regulr expression r strts with the singleton set {r}. More precisely, prtil derivtives pder :: α α rexp (α rexp) set re defined recursively s follows: pder _ 0 = {} pder _ 1 = {} pder (A x) = if x = then {1} else {} pder (r + s) = pder r pder s pder (r s) = if nullle r then (pder r s) pder s else pder r s pder (r ) = pder r r

8 Aove, R s is used s shorthnd nottion for {r s r R}. The definition yields the chrcteristic property of prtil derivtives y induction on r: D (L r) = s pder r L s Following this chrcteristic property, we cn interpret the locle rexpdfa. The utomton constructed y P for our running exmple is shown in Fig. 4. interprettion P: rexpdfa where ι r = {r} δ R = r R pder r o R = r R. nullle r L R = r R L r { } {(1 ) } {1} {} Fig. 4: Prtil derivtive utomton for The ssumptions of rexpda (inherited y rexpdfa) re esy to dischrge. Just s for Brzozowski derivtives, only the proof of finiteness of the rechle stte spce P.fin poses chllenge. We were le to reuse the proof y Wu et l. [27] who show finiteness when proving one direction of the Myhill-Nerode theorem. Compred with the proof of D.fin, the forml resoning out prtil derivtives ppers to e more succinct. There is direct connection etween pder nd der tht seems not to hve een covered in the literture. It is est expressed in terms of recursive function pset :: αrexp (α rexp) set tht trnsltes derivtives to prtil derivtives: pset (der r) = pder r. pset 0 = {} pset (r + s) = pset r pset s pset 1 = {1} pset (r s) = pset r s pset (A x) = {A x} pset (r ) = {r } A finite set R of regulr expressions cn e represented uniquely y single regulr expression R, sum ordered w.r.t.. Hence, we hve pset (der r) = pder r, mening tht we cn devise normliztion function r = pset r tht llows us to simulte prtil derivtives while operting on plin regulr expressions. Alterntively, _ cn e defined using smrt constructors (with sequentilly mtched equtions): 0 = 0 1 = 1 A = A r + s = r s r s = r s r = r 0 r = 0 (r + s) t = (r s) (s t) r s = s t 0 r = r r 0 = r (r + s) t = r (s t) r (s + t) = if r = s then s + t else if r s then r + (s + t) else s + (r t) r s = if r = s then r else if r s then r + s else s + r This definition llows to contrst the implicit quotienting performed y prtil derivtives with the qoutienting modulo ACI equivlence ( ). They turn out to e incomprle: _ does not simplify the second rgument of conctention nd the rgument of itertion, ut erses 0s nd uses left distriutivity. Finlly, we otin lst derivtive-sed interprettion using the chrcteristic property der r = (pder r) nd P.fin to dischrge the finiteness ssumption fin.

9 interprettion PD: rexpdfa where ι r = r δ r = der r o r = nullle r L r = L r Whenever P yields n utomton for r with sttes leled with finite sets of regulr expressions X i, PD constructs structurlly the sme utomton for r leled with X i. 5 Mrked Regulr Expressions One of the oldest methods for converting regulr expression into n utomton is sed on the ide of identifying the sttes of the utomton with positions in the regulr expression. Both McNughton nd Ymd [20] nd Glushkov [13] mrk the toms in regulr expression with numers to identify positions uniquely. In this section, we formlize two recent reincrntions of this pproch due to Fischer et l. [11] nd Asperti [3]. They re sed on the reliztion tht in functionl progrmming setting, it is most convenient to represent positions in regulr expression y mrking some of its toms. First we define n infrstructure for working with mrked regulr expressions. Then we define nd relte oth reincrntions in terms of this infrstructure. Mrked regulr expressions re formlized y the following type synonym (where the vlue True denotes mrked tom) α mrexp = (ool α)rexp We convert esily etween rexp nd mrexp with the help of mp_rexp, the mp function on regulr expressions: strip = mp_rexp snd emtpy_ mrexp = mp_rexp (λr. (Flse, r)) The lnguge L m :: α mrexp αlng of mrked regulr expression is the set of words tht strt t some mrked tom: L m 0 = {} L m 1 = {} L m (A (m, )) = if m then {[]} else {} L m (r + s) = L m r L m s L m (r s) = (L m r L (strip s)) L m s L m (r ) = L m r L (strip r) The function finl :: α mrexp ool tests if some tom t the end of given regulr expression is mrked: finl 0 = Flse finl 1 = Flse finl (A (m, )) = m finl (r + s) = (finl r finl s) finl (r s) = (finl s nullle s finl r) finl (r ) = finl r

10 Mrks re moved round regulr expression y two opertions. The function red r unmrks ll toms in r except : red :: α α mrexp α mrexp red = mp_rexp (λ(m, x). (m = x, x)) Its chrcteristic lemm is tht it restricts L m r to words whose hed is : L m (red r) = {w L m r w [] hd w = } The function follow m r moves ll mrks in r to the next tom, much like n ε-closure; the mrk m is pushed in from the left: follow :: ool α mrexp α mrexp follow m 0 = 0 follow m 1 = 1 follow m (A (_, )) = A (m, ) follow m (r + s) = follow m r + follow m s follow m (r s) = follow m r follow (finl r m nullle r) s follow m (r ) = (follow (finl r m) r) The chrcteristic lemm out follow shows tht the mrks re moved forwrd, therey chopping off the first letter (in the generted lnguge), nd tht the prmeter m indictes whether every first tom should e mrked: L m (follow m r) = {tl w w L m r} (if m then L (strip r) else {}) {[]} 5.1 Mrk After Atom In the work of McNughton-Ymd-Glushkov, the mrk indictes which tom hs just een red, i.e. the mrk is locted fter the tom. Therefore the initil stte is specil ecuse nothing hs een red yet. Thus we express the sttes of the utomton s pir of oolen (True mens tht nothing hs een red yet) nd mrked regulr expression. The oolen cn e viewed s mrk in front of the utomton. (Alterntively, one could work with n explicit strt symol in front of the regulr expression.) We interpret the locle rexpdfa s follows: interprettion A: rexpdfa where ι r = (True, emtpy_ mrexp r) δ (m, r) = (Flse, red (follow m r)) o (m, r) = (finl r m nullle r) L (m, r) = L m (follow m r) (if o (m, r) then {[]} else {}) The definition of δ expresses tht we first uild the ε-closure strting from the mrked toms (vi follow) nd then red the next tom. With the chrcteristic lemms out red nd follow (nd few uxiliry lemms), the locle ssumptions re esily proved. This yields our first version of utomt sed on mrked regulr expressions. Finiteness of the rechle prt of the stte spce is proved vi the lemm fold δ w (ι r) {True, Flse} mrexps r

11 where mrexps :: αrexp (α mrexp)set mps regulr expression to the finite set of ll its mrked vrints, i.e. mrexps r = {r strip r = r}; its ctul recursive definition is strightforwrd nd omitted. Now we tke closer look t the work of Fischer et l. [11], which inspired the preceding formliztion. They present numer of (not formlly verified) mtching lgorithms on mrked regulr expressions in Hskell tht follow McNughton-Ymd- Glushkov. This is their sic trnsition function: shift :: ool α mrexp α α mrexp shift _ 0 _ = 0 shift _ 1 _ = 1 shift m (A (_, x)) c = A (m (x = c), x) shift m (r + s) c = shift m r c + shift m s c shift m (r s) c = shift m r c shift (finl r m nullle r) s c shift m (r ) c = (shift (finl r m) r c) A simple induction proves tht their shift is our δ: shift m r x = red x (follow m r) Thus we hve verified their shift function. Fischer et l. optimize shift further, which is still qudrtic due to the clls of the recursive functions finl nd nullle. They simply cche the vlues of finl nd nullle t ll nodes of regulr expression y dding dditionl fields to ech constructor. We hve verified this optimiztion step s well, yielding nother interprettion A 2 (omitted here). 5.2 Mrk Before Atom Insted of imgining the mrk to e fter n tom, it cn lso e viewed to e in front of it, i.e. it mrks possile next toms. This is somewht dul to the McNughton-Ymd- Glushkov construction. It leds to the following interprettion of the rexpda locle: interprettion B: rexpdfa where ι r = (follow True (emtpy_ mrexp r), nullle r) δ (r, m) = let r = red r in (follow Flse r, finl r ) o (r, m) = m L (r, m) = L m r (if m then {[]} else {}) The definition of δ expresses tht we first red n tom nd then uild the ε-closure. The ssumptions of rexpda nd rexpdfa re proved esily just like in the previous interprettion with mrked regulr expressions. The interesting point is tht this hppens to e the lgorithm formlized y Asperti [3]. Although he sys tht he hs formlized McNughton-Ymd, he ctully formlized the dul lgorithm. This is not esy to see ecuse Asperti s formliztion is considerly more involved thn ours, with mny uxiliry functions. Strictly speking, his lgorithm is vrition of ours tht produces the sme utomt. The complete proof of this fct cn e found elsewhere [16]. Becuse of the size of Asperti s formliztion, there is not enough spce here to give the detiled equivlence proof. However, we cn tke step towrds his formultion nd merge follow nd red into one function move :: α α mrexp ool α mrexp, the nlogue of his homonymous function:

12 ( ) ( ) ( ) ( ) ( ) ( ) Fig. 5: Mrked regulr expression utomt (A left, B right) for move _ 0 _ = 0 move _ 1 _ = 1 move c (A (_, x)) m = A (m, x) move c (r + s) m = move c r m + move c s m move c (r s) m = move c r m move c s (finl1 r c m nullle r) move c (r ) m = (move c r (finl1 r c m)) where finl1 is n uxiliry recursive function (not shown here) with the chrcteristic property tht finl1 r c = finl (red c r). A simple induction proves tht move comines follow nd red s in δ: move c r m = follow m (red c r) The function move hs qudrtic complexity for the sme reson s shift. Unfortuntely, it cnnot e mde liner with the sme ese s for shift. The prolem is tht we need to cche the vlue of finl1 r c in the previous step, efore we know c. We solve this y cching the set of ll letters c tht mke finl1 r c true. In the worst cse, the whole lphet must e stored in certin inner nodes. However, for n lphet of fixed size this gurntees liner time complexity. This optimiztion constitutes lst interprettion B 2. Even for fixed lphet, Asperti s move hs qudrtic complexity when fced with tower of strs: ech recursive cll of move cn trigger cll of function eclose, which hs liner complexity. Asperti imed for compct proofs, not mximl efficiency. 5.3 Comprison The two constructions my look similr, ut they do not produce isomorphic utomt. Considering our running exmple, we disply the mrk y efore or fter the tom. The two resulting utomt re shown in Fig. 5. There re specil sttes tht cnnot e denoted y mrking toms only: r in A s utomton is the completely unmrked regulr expression tht is the initil stte nd r in B s utomton is finl stte. It turns out tht the efore utomton is homomorphic imge of the fter utomton. To verify this we specify the homomorphism ϕ(m, r) = (follow m r, A.o (m, r)) nd prove tht it preserves initil sttes nd commutes with the trnsition function: ϕ(a.ι r) = B.ι r ϕ(a.δ s) = B.δ (ϕ s) ϕ(fold A.δ w s) = fold B.δ w (ϕ s) A direct consequence is tht Asperti s efore construction lwys genertes utomt with t most s mny sttes s the McNughton-Ymd-Glushkov construction. Formlly, in the context of locle rexpda we hve defined n executle computtion of the rechle stte spce {fold δ w (ι r) w (set s) } of the utomton: rechle s r = snd (the (rtc (λ_. True) (λ s. mp (λ. δ s) s) (ι r)))

13 where r is the initil regulr expression, s is the lphet, nd the (Some x) = x. Theorem 5. B.rechle s r A.rechle s r where _ is the crdinlity of set. In erly drfts of this pper, we only conjectured the ove sttement nd unsuccessfully tried to refute it with Iselle s Quickcheck fcility [8]. Lter, Helmut Seidl hs communicted n informl proof using the ove homomorphism to us. Let us revite the sttement of Thm. 5 to n n. One my think tht n is only slightly lrger thn n, ut it seems tht n nd n re more thn constnt summnd prt: for two-element lphet Quickcheck could refute n n +k even for k = Empiricl Comprison We compre the efficiency w.r.t. oth mtching nd deciding equivlence of the Stndrd ML code generted from eight descried interprettions: -normlized derivtives (D), -normlized derivtives (N), prtil derivtives (P), derivtives simulting prtil derivtives (PD), mrk fter tom (A), mrk fter tom with cching (A 2 ), mrk efore tom (B), nd mrk efore tom with cching (B 2 ). The interprettion using the quotient type for derivtives (D ) is not in this list, s it is clerly superseded y D. The results of the evlution, performed on n Intel Core i7-2760qm mchine with 8GB of RAM, re shown in Fig. 6. Solid lines depict the four derivtive-sed lgorithms. Dshed lines re used for the lgorithms sed on mrked regulr expressions. The first two tests, MATCH-R nd MATCH-L, mesure the time required to mtch the word n ginst the regulr expression ( + 1) n n stndrd enchmrk lso used y Fischer et l. [11]. The difference etween the two tests is the definition of r n. MATCH-R defines it s the n-fold conctention ssocited to the right: r 4 = r (r (r r)), wheres MATCH-L ssocites to the left: r 4 = ((r r) r) r. In oth tests, mrked regulr expressions outperform derivtives y fr. The normliztion performed y the derivtive-sed pproches (required to otin finite numer of sttes for the equivlence check) decelertes the computtion of the next stte. Mrked regulr expressions enefit from fst next stte computtion. The test MATCH-L exhiits the qudrtic nture of the unoptimized mtchers A nd B (their curves re lmost identicl nd therefore hrd to distinguish in Fig. 6). In contrst, A 2 nd B 2 perform eqully well in oth tests, A 2 eing pproximtely 1.5 times fster due to lighter cche nnottions. The next test goes ck to Antimirov [1]: We mesure the time (with timeout of ten seconds) for proving the equivlence of nd ( n 1 ) ( n ). Agin two tests, EQ-R nd EQ-L, distinguish the ssocitivity of conctention in r n. Here, the derivtive-sed equivlence checkers (except for D) perform etter then the ones sed on mrked regulr expressions. In prticulr, oth version of prtil derivtives, P nd PD, outperform N since this exmple ws crfted y Antimirov to demonstrte the strength of prtil derivtives, this is not wholly unexpected. Compring EQ-R nd EQ-L, the ssocitivity rely influences the runtime. Finlly, to void is towrds prticulr lgorithm, we hve devised the rndomized test EQ-RND. There we mesure the verge time (with timeout of ten seconds) to prove the equivlence of r with itself for 100 rndomly generted expressions with n inner nodes (+,, or ). Proving r r is of course trivil tsk, ut our lgorithms do not

14 4 (MATCH-R) 4 (MATCH-L) Time (s) 2 Time (s) n n (EQ-R) (EQ-L) 8 8 Time (s) n Time (s) n Time (s) 4 2 (EQ-RND) n D N P PD B B 2 A A 2 Fig. 6: Evlution results stop the explortion when the stte of the product utomton is pir of two equl sttes. This optimiztion, which is must for ny prcticl lgorithm, is the first step towrds the rewrding usge of isimultion up to equivlence (or even up to congruence) [5]. Without ny such optimiztion, the tsk of proving r r mounts to enumerting ll derivtives of r, which is exctly wht we wnt to compre. To generte rndom regulr expression we use the infrstructure of SpecCheck [25] Quickcheck clone for Iselle/ML. For computing the verge, timeout counts s 10 second (lthough the ctul computtion would likely hve tken longer) n pproximtion tht skews the curves to converge to the mrgin of 10 seconds. We stopped mesuring method for incresing n when the verge pproched 5 seconds. The results of EQ-RND re summrized s follows: D N P,PD A,A 2,B,B 2, where X Y mens tht Y is n order of mgnitude fster thn X. The lgorithm P is noticely slower thn PD voiding sets reduces the overhed. Among A, A 2, B, B 2, Asperti s unoptimized lgorithm B performs est y nrrow mrgin. Regulr expressions where the cching overhed pys off re rre nd therefore not visile in the rndomized test results. The sme holds for expressions where B produces much smller utomt thn A (e.g. the counterexmple to n n from Susect. 5.3). Our evlution shows tht A 2 is the est choice for mtching. For equivlence checking, the winner is not s cler cut: B (especilly when pplied to normlized input to void qudrtic runtime without cching) nd PD seem to e the est choices.

15 7 Extensions Brzozowski s derivtives re esily extendle to regulr expressions intersection nd negtion indeed Brzozowski performed the extension right from the strt [7]. The numer of such extended derivtives is still finite when quotiented modulo ACI. We [26] hve recently further extended derivtives to regulr expressions extended with projection, otining verified decision procedures for the equivlence of those extended regulr expressions nd for mondic second-order logics over finite words. The closure computtion nd its correctness proof follow Kruss nd Nipkow [19]. Extending prtil derivtives with intersection nd negtion is more involved [9]. An dditionl lyer of sets must e used for intersections, i.e. the sttes of our utomton would then e sets of sets of regulr expressions. In Sect. 6, we hve seen tht lredy one lyer of sets incurs some overhed. Hence, the view on prtil derivtives s derivtives followed y some normliztion is expected to e even more profitle for the extension. The extension of prtil derivtives with projection is n esy exercise. It is uncler how to extend mrked regulr expressions to hndle negtion nd intersection. The numer of possile mrkings for regulr expression of lphetic width n is 2 n. However, there exist regulr expressions of lphetic width n using intersection, whose miniml utomt hve 2 2n sttes [12]. 8 Conclusion We hve shown tht ll the previously pulished verified decision procedures for equivlence of regulr expressions tht operte on regulr expressions directly cn ll e expressed s instnces of generic utomton-inspired frmework. The correctness proofs decompose into generic prt tht is proved once nd for ll in the frmework nd few specific properties tht need to e proved for ech instnce. The frmework cters for meningful comprison of the performnce of the vrious instnces. Mrked regulr expressions re superior on verge ut prtil derivtives cn outperform them in specific cses. The Iselle theories re ville online [23]. Acknowledgment. We thnk Andre Asperti nd Sestin Fischer for commenting on fine points of their work nd Helmut Seidl for contriuting n informl proof of Thm. 5. Jsmin Blnchette, Andrei Popescu nd three nonymous reviewers helped to improve the presenttion through numerous suggestions. The second uthor is supported y the doctorte progrm 1480 (PUMA) of the Deutsche Forschungsgemeinschft (DFG). References 1. Antimirov, V.: Prtil derivtives of regulr expressions nd finite utomt constructions. In: Myr, E.W., Puech, C. (eds.) STACS 95. LNCS, vol. 900, pp Springer (1995) 2. Antimirov, V.: Prtil derivtives of regulr expressions nd finite utomton constructions. Theor. Comput. Sci. 155(2), (1996) 3. Asperti, A.: A compct proof of decidility for regulr expression equivlence. In: Beringer, L., Felty, A. (eds.) ITP LNCS, vol. 7406, pp Springer (2012) 4. Bllrin, C.: Interprettion of locles in Iselle: Theories nd proof contexts. In: Borwein, J.M., Frmer, W.M. (eds.) MKM LNCS, vol. 4108, pp Springer (2006)

16 5. Bonchi, F., Pous, D.: Checking NFA equivlence with isimultions up to congruence. In: Gicozzi, R., Cousot, R. (eds.) POPL pp ACM (2013) 6. Brint, T., Pous, D.: An efficient Coq tctic for deciding Kleene lgers. In: Kufmnn, M., Pulson, L. (eds.) ITP LNCS, vol. 6172, pp Springer (2010) 7. Brzozowski, J.A.: Derivtives of regulr expressions. J. ACM 11(4), (1964) 8. Bulwhn, L.: The new Quickcheck for Iselle: Rndom, exhustive nd symolic testing under one roof. In: Hwlitzel, C., Miller, D. (eds.) CPP LNCS, vol. 7679, pp Springer (2012) 9. Cron, P., Chmprnud, J.M., Mignot, L.: Prtil derivtives of n extended regulr expression. In: Dediu, A.H., Ineng, S., Mrtín-Vide, C. (eds.) LATA LNCS, vol. 6638, pp Springer (2011) 10. Coqund, T., Siles, V.: A decision procedure for regulr expression equivlence in type theory. In: Jounnud, J.P., Sho, Z. (eds.) CPP LNCS, vol. 7086, pp Springer (2011) 11. Fischer, S., Huch, F., Wilke, T.: A ply on regulr expressions: functionl perl. In: Hudk, P., Weirich, S. (eds.) ICFP pp ACM (2010) 12. Gelde, W., Neven, F.: Succinctness of the complement nd intersection of regulr expressions. ACM Trns. Comput. Log. 13(1), 4:1 19 (2012) 13. Glushkov, V.M.: The strct theory of utomt. Russin Mth. Surveys 16, 1 53 (1961) 14. Hftmnn, F., Kruss, A., Kunčr, O., Nipkow, T.: Dt refinement in Iselle/HOL. In: Blzy, S., Pulin-Mohring, C., Pichrdie, D. (eds.) ITP LNCS, vol. 7998, pp Springer (2013) 15. Hftmnn, F., Nipkow, T.: Code genertion vi higher-order rewrite systems. In: Blume, M., Koyshi, N., Vidl, G. (eds.) FLOPS LNCS, vol. 6009, pp Springer (2010) 16. Hsleck, M.: Verified Decision Procedures for the Equivlence of Regulr Expressions. B.Sc. thesis, Deprtment of Informtics, Technische Universität München (2013) 17. Hufmnn, B., Kunčr, O.: Lifting nd Trnsfer: A modulr design for quotients in Iselle/HOL. In: Gonthier, G., Norrish, M. (eds.) CPP LNCS, vol. 8307, pp Springer (2013) 18. Kliszyk, C., Urn, C.: Quotients revisited for Iselle/HOL. In: Chu, W.C., Wong, W.E., Plkl, M.J., Hung, C.C. (eds.) SAC pp ACM (2011) 19. Kruss, A., Nipkow, T.: Proof perl: Regulr expression equivlence nd reltion lger. J. Automted Resoning 49, (2012), pulished online Mrch McNughton, R., Ymd, H.: Regulr expressions nd finite stte grphs for utomt. IRE Trns. on Electronic Comput EC-9, (1960) 21. Moreir, N., Pereir, D., de Sous, S.M.: Deciding regulr expressions (in-)equivlence in Coq. In: Khl, W., Griffin, T. (eds.) RAMiCS LNCS, vol. 7560, pp Springer (2012) 22. Nipkow, T., Klein, G.: Concrete Semntics. A Proof Assistnt Approch. Springer (to pper), Nipkow, T., Trytel, D.: Regulr expression equivlence. Archive of Forml Proofs (2014), Forml proof development 24. Rutten, J.J.M.M.: Automt nd coinduction (n exercise in colger). In: Sngiorgi, D., de Simone, R. (eds.) CONCUR LNCS, vol. 1466, pp Springer (1998) 25. Schffroth, N.: A Specifiction-sed Testing Tool for Iselle s ML Environment. B.Sc. thesis, Deprtment of Informtics, Technische Universität München (2013) 26. Trytel, D., Nipkow, T.: Verified decision procedures for MSO on words sed on derivtives of regulr expressions. In: Morrisett, G., Uustlu, T. (eds.) ICFP pp (2013) 27. Wu, C., Zhng, X., Urn, C.: A formlistion of the Myhill-Nerode theorem sed on regulr expressions. J. Automted Resoning, 52, (2014)

Convert the NFA into DFA

Convert the NFA into DFA Convert the NF into F For ech NF we cn find F ccepting the sme lnguge. The numer of sttes of the F could e exponentil in the numer of sttes of the NF, ut in prctice this worst cse occurs rrely. lgorithm:

More information

Coalgebra, Lecture 15: Equations for Deterministic Automata

Coalgebra, Lecture 15: Equations for Deterministic Automata Colger, Lecture 15: Equtions for Deterministic Automt Julin Slmnc (nd Jurrin Rot) Decemer 19, 2016 In this lecture, we will study the concept of equtions for deterministic utomt. The notes re self contined

More information

CMPSCI 250: Introduction to Computation. Lecture #31: What DFA s Can and Can t Do David Mix Barrington 9 April 2014

CMPSCI 250: Introduction to Computation. Lecture #31: What DFA s Can and Can t Do David Mix Barrington 9 April 2014 CMPSCI 250: Introduction to Computtion Lecture #31: Wht DFA s Cn nd Cn t Do Dvid Mix Brrington 9 April 2014 Wht DFA s Cn nd Cn t Do Deterministic Finite Automt Forml Definition of DFA s Exmples of DFA

More information

1 Nondeterministic Finite Automata

1 Nondeterministic Finite Automata 1 Nondeterministic Finite Automt Suppose in life, whenever you hd choice, you could try oth possiilities nd live your life. At the end, you would go ck nd choose the one tht worked out the est. Then you

More information

Lecture 08: Feb. 08, 2019

Lecture 08: Feb. 08, 2019 4CS4-6:Theory of Computtion(Closure on Reg. Lngs., regex to NDFA, DFA to regex) Prof. K.R. Chowdhry Lecture 08: Fe. 08, 2019 : Professor of CS Disclimer: These notes hve not een sujected to the usul scrutiny

More information

Minimal DFA. minimal DFA for L starting from any other

Minimal DFA. minimal DFA for L starting from any other Miniml DFA Among the mny DFAs ccepting the sme regulr lnguge L, there is exctly one (up to renming of sttes) which hs the smllest possile numer of sttes. Moreover, it is possile to otin tht miniml DFA

More information

Formal Languages and Automata

Formal Languages and Automata Moile Computing nd Softwre Engineering p. 1/5 Forml Lnguges nd Automt Chpter 2 Finite Automt Chun-Ming Liu cmliu@csie.ntut.edu.tw Deprtment of Computer Science nd Informtion Engineering Ntionl Tipei University

More information

CS415 Compilers. Lexical Analysis and. These slides are based on slides copyrighted by Keith Cooper, Ken Kennedy & Linda Torczon at Rice University

CS415 Compilers. Lexical Analysis and. These slides are based on slides copyrighted by Keith Cooper, Ken Kennedy & Linda Torczon at Rice University CS415 Compilers Lexicl Anlysis nd These slides re sed on slides copyrighted y Keith Cooper, Ken Kennedy & Lind Torczon t Rice University First Progrmming Project Instruction Scheduling Project hs een posted

More information

Assignment 1 Automata, Languages, and Computability. 1 Finite State Automata and Regular Languages

Assignment 1 Automata, Languages, and Computability. 1 Finite State Automata and Regular Languages Deprtment of Computer Science, Austrlin Ntionl University COMP2600 Forml Methods for Softwre Engineering Semester 2, 206 Assignment Automt, Lnguges, nd Computility Smple Solutions Finite Stte Automt nd

More information

Lecture 9: LTL and Büchi Automata

Lecture 9: LTL and Büchi Automata Lecture 9: LTL nd Büchi Automt 1 LTL Property Ptterns Quite often the requirements of system follow some simple ptterns. Sometimes we wnt to specify tht property should only hold in certin context, clled

More information

The University of Nottingham SCHOOL OF COMPUTER SCIENCE A LEVEL 2 MODULE, SPRING SEMESTER LANGUAGES AND COMPUTATION ANSWERS

The University of Nottingham SCHOOL OF COMPUTER SCIENCE A LEVEL 2 MODULE, SPRING SEMESTER LANGUAGES AND COMPUTATION ANSWERS The University of Nottinghm SCHOOL OF COMPUTER SCIENCE LEVEL 2 MODULE, SPRING SEMESTER 2016 2017 LNGUGES ND COMPUTTION NSWERS Time llowed TWO hours Cndidtes my complete the front cover of their nswer ook

More information

1. For each of the following theorems, give a two or three sentence sketch of how the proof goes or why it is not true.

1. For each of the following theorems, give a two or three sentence sketch of how the proof goes or why it is not true. York University CSE 2 Unit 3. DFA Clsses Converting etween DFA, NFA, Regulr Expressions, nd Extended Regulr Expressions Instructor: Jeff Edmonds Don t chet y looking t these nswers premturely.. For ech

More information

Java II Finite Automata I

Java II Finite Automata I Jv II Finite Automt I Bernd Kiefer Bernd.Kiefer@dfki.de Deutsches Forschungszentrum für künstliche Intelligenz Finite Automt I p.1/13 Processing Regulr Expressions We lredy lerned out Jv s regulr expression

More information

Finite Automata-cont d

Finite Automata-cont d Automt Theory nd Forml Lnguges Professor Leslie Lnder Lecture # 6 Finite Automt-cont d The Pumping Lemm WEB SITE: http://ingwe.inghmton.edu/ ~lnder/cs573.html Septemer 18, 2000 Exmple 1 Consider L = {ww

More information

Intermediate Math Circles Wednesday, November 14, 2018 Finite Automata II. Nickolas Rollick a b b. a b 4

Intermediate Math Circles Wednesday, November 14, 2018 Finite Automata II. Nickolas Rollick a b b. a b 4 Intermedite Mth Circles Wednesdy, Novemer 14, 2018 Finite Automt II Nickols Rollick nrollick@uwterloo.c Regulr Lnguges Lst time, we were introduced to the ide of DFA (deterministic finite utomton), one

More information

Finite Automata Theory and Formal Languages TMV027/DIT321 LP4 2018

Finite Automata Theory and Formal Languages TMV027/DIT321 LP4 2018 Finite Automt Theory nd Forml Lnguges TMV027/DIT321 LP4 2018 Lecture 10 An Bove April 23rd 2018 Recp: Regulr Lnguges We cn convert between FA nd RE; Hence both FA nd RE ccept/generte regulr lnguges; More

More information

Chapter Five: Nondeterministic Finite Automata. Formal Language, chapter 5, slide 1

Chapter Five: Nondeterministic Finite Automata. Formal Language, chapter 5, slide 1 Chpter Five: Nondeterministic Finite Automt Forml Lnguge, chpter 5, slide 1 1 A DFA hs exctly one trnsition from every stte on every symol in the lphet. By relxing this requirement we get relted ut more

More information

Nondeterminism and Nodeterministic Automata

Nondeterminism and Nodeterministic Automata Nondeterminism nd Nodeterministic Automt 61 Nondeterminism nd Nondeterministic Automt The computtionl mchine models tht we lerned in the clss re deterministic in the sense tht the next move is uniquely

More information

CMSC 330: Organization of Programming Languages

CMSC 330: Organization of Programming Languages CMSC 330: Orgniztion of Progrmming Lnguges Finite Automt 2 CMSC 330 1 Types of Finite Automt Deterministic Finite Automt (DFA) Exctly one sequence of steps for ech string All exmples so fr Nondeterministic

More information

Lecture 09: Myhill-Nerode Theorem

Lecture 09: Myhill-Nerode Theorem CS 373: Theory of Computtion Mdhusudn Prthsrthy Lecture 09: Myhill-Nerode Theorem 16 Ferury 2010 In this lecture, we will see tht every lnguge hs unique miniml DFA We will see this fct from two perspectives

More information

Converting Regular Expressions to Discrete Finite Automata: A Tutorial

Converting Regular Expressions to Discrete Finite Automata: A Tutorial Converting Regulr Expressions to Discrete Finite Automt: A Tutoril Dvid Christinsen 2013-01-03 This is tutoril on how to convert regulr expressions to nondeterministic finite utomt (NFA) nd how to convert

More information

Model Reduction of Finite State Machines by Contraction

Model Reduction of Finite State Machines by Contraction Model Reduction of Finite Stte Mchines y Contrction Alessndro Giu Dip. di Ingegneri Elettric ed Elettronic, Università di Cgliri, Pizz d Armi, 09123 Cgliri, Itly Phone: +39-070-675-5892 Fx: +39-070-675-5900

More information

Types of Finite Automata. CMSC 330: Organization of Programming Languages. Comparing DFAs and NFAs. Comparing DFAs and NFAs (cont.) Finite Automata 2

Types of Finite Automata. CMSC 330: Organization of Programming Languages. Comparing DFAs and NFAs. Comparing DFAs and NFAs (cont.) Finite Automata 2 CMSC 330: Orgniztion of Progrmming Lnguges Finite Automt 2 Types of Finite Automt Deterministic Finite Automt () Exctly one sequence of steps for ech string All exmples so fr Nondeterministic Finite Automt

More information

Lecture 3: Equivalence Relations

Lecture 3: Equivalence Relations Mthcmp Crsh Course Instructor: Pdric Brtlett Lecture 3: Equivlence Reltions Week 1 Mthcmp 2014 In our lst three tlks of this clss, we shift the focus of our tlks from proof techniques to proof concepts

More information

Formal languages, automata, and theory of computation

Formal languages, automata, and theory of computation Mälrdlen University TEN1 DVA337 2015 School of Innovtion, Design nd Engineering Forml lnguges, utomt, nd theory of computtion Thursdy, Novemer 5, 14:10-18:30 Techer: Dniel Hedin, phone 021-107052 The exm

More information

Automata Theory 101. Introduction. Outline. Introduction Finite Automata Regular Expressions ω-automata. Ralf Huuck.

Automata Theory 101. Introduction. Outline. Introduction Finite Automata Regular Expressions ω-automata. Ralf Huuck. Outline Automt Theory 101 Rlf Huuck Introduction Finite Automt Regulr Expressions ω-automt Session 1 2006 Rlf Huuck 1 Session 1 2006 Rlf Huuck 2 Acknowledgement Some slides re sed on Wolfgng Thoms excellent

More information

Types of Finite Automata. CMSC 330: Organization of Programming Languages. Comparing DFAs and NFAs. NFA for (a b)*abb.

Types of Finite Automata. CMSC 330: Organization of Programming Languages. Comparing DFAs and NFAs. NFA for (a b)*abb. CMSC 330: Orgniztion of Progrmming Lnguges Finite Automt 2 Types of Finite Automt Deterministic Finite Automt () Exctly one sequence of steps for ech string All exmples so fr Nondeterministic Finite Automt

More information

Designing finite automata II

Designing finite automata II Designing finite utomt II Prolem: Design DFA A such tht L(A) consists of ll strings of nd which re of length 3n, for n = 0, 1, 2, (1) Determine wht to rememer out the input string Assign stte to ech of

More information

Parse trees, ambiguity, and Chomsky normal form

Parse trees, ambiguity, and Chomsky normal form Prse trees, miguity, nd Chomsky norml form In this lecture we will discuss few importnt notions connected with contextfree grmmrs, including prse trees, miguity, nd specil form for context-free grmmrs

More information

AUTOMATA AND LANGUAGES. Definition 1.5: Finite Automaton

AUTOMATA AND LANGUAGES. Definition 1.5: Finite Automaton 25. Finite Automt AUTOMATA AND LANGUAGES A system of computtion tht only hs finite numer of possile sttes cn e modeled using finite utomton A finite utomton is often illustrted s stte digrm d d d. d q

More information

Chapter 2 Finite Automata

Chapter 2 Finite Automata Chpter 2 Finite Automt 28 2.1 Introduction Finite utomt: first model of the notion of effective procedure. (They lso hve mny other pplictions). The concept of finite utomton cn e derived y exmining wht

More information

5. (±±) Λ = fw j w is string of even lengthg [ 00 = f11,00g 7. (11 [ 00)± Λ = fw j w egins with either 11 or 00g 8. (0 [ ffl)1 Λ = 01 Λ [ 1 Λ 9.

5. (±±) Λ = fw j w is string of even lengthg [ 00 = f11,00g 7. (11 [ 00)± Λ = fw j w egins with either 11 or 00g 8. (0 [ ffl)1 Λ = 01 Λ [ 1 Λ 9. Regulr Expressions, Pumping Lemm, Right Liner Grmmrs Ling 106 Mrch 25, 2002 1 Regulr Expressions A regulr expression descries or genertes lnguge: it is kind of shorthnd for listing the memers of lnguge.

More information

CS 275 Automata and Formal Language Theory

CS 275 Automata and Formal Language Theory CS 275 Automt nd Forml Lnguge Theory Course Notes Prt II: The Recognition Problem (II) Chpter II.6.: Push Down Automt Remrk: This mteril is no longer tught nd not directly exm relevnt Anton Setzer (Bsed

More information

Tutorial Automata and formal Languages

Tutorial Automata and formal Languages Tutoril Automt nd forml Lnguges Notes for to the tutoril in the summer term 2017 Sestin Küpper, Christine Mik 8. August 2017 1 Introduction: Nottions nd sic Definitions At the eginning of the tutoril we

More information

Name Ima Sample ASU ID

Name Ima Sample ASU ID Nme Im Smple ASU ID 2468024680 CSE 355 Test 1, Fll 2016 30 Septemer 2016, 8:35-9:25.m., LSA 191 Regrding of Midterms If you elieve tht your grde hs not een dded up correctly, return the entire pper to

More information

Bases for Vector Spaces

Bases for Vector Spaces Bses for Vector Spces 2-26-25 A set is independent if, roughly speking, there is no redundncy in the set: You cn t uild ny vector in the set s liner comintion of the others A set spns if you cn uild everything

More information

More on automata. Michael George. March 24 April 7, 2014

More on automata. Michael George. March 24 April 7, 2014 More on utomt Michel George Mrch 24 April 7, 2014 1 Automt constructions Now tht we hve forml model of mchine, it is useful to mke some generl constructions. 1.1 DFA Union / Product construction Suppose

More information

CSCI 340: Computational Models. Kleene s Theorem. Department of Computer Science

CSCI 340: Computational Models. Kleene s Theorem. Department of Computer Science CSCI 340: Computtionl Models Kleene s Theorem Chpter 7 Deprtment of Computer Science Unifiction In 1954, Kleene presented (nd proved) theorem which (in our version) sttes tht if lnguge cn e defined y ny

More information

Regular expressions, Finite Automata, transition graphs are all the same!!

Regular expressions, Finite Automata, transition graphs are all the same!! CSI 3104 /Winter 2011: Introduction to Forml Lnguges Chpter 7: Kleene s Theorem Chpter 7: Kleene s Theorem Regulr expressions, Finite Automt, trnsition grphs re ll the sme!! Dr. Neji Zgui CSI3104-W11 1

More information

Compiler Design. Fall Lexical Analysis. Sample Exercises and Solutions. Prof. Pedro C. Diniz

Compiler Design. Fall Lexical Analysis. Sample Exercises and Solutions. Prof. Pedro C. Diniz University of Southern Cliforni Computer Science Deprtment Compiler Design Fll Lexicl Anlysis Smple Exercises nd Solutions Prof. Pedro C. Diniz USC / Informtion Sciences Institute 4676 Admirlty Wy, Suite

More information

CS103B Handout 18 Winter 2007 February 28, 2007 Finite Automata

CS103B Handout 18 Winter 2007 February 28, 2007 Finite Automata CS103B ndout 18 Winter 2007 Ferury 28, 2007 Finite Automt Initil text y Mggie Johnson. Introduction Severl childrens gmes fit the following description: Pieces re set up on plying ord; dice re thrown or

More information

State Minimization for DFAs

State Minimization for DFAs Stte Minimiztion for DFAs Red K & S 2.7 Do Homework 10. Consider: Stte Minimiztion 4 5 Is this miniml mchine? Step (1): Get rid of unrechle sttes. Stte Minimiztion 6, Stte is unrechle. Step (2): Get rid

More information

3 Regular expressions

3 Regular expressions 3 Regulr expressions Given n lphet Σ lnguge is set of words L Σ. So fr we were le to descrie lnguges either y using set theory (i.e. enumertion or comprehension) or y n utomton. In this section we shll

More information

Harvard University Computer Science 121 Midterm October 23, 2012

Harvard University Computer Science 121 Midterm October 23, 2012 Hrvrd University Computer Science 121 Midterm Octoer 23, 2012 This is closed-ook exmintion. You my use ny result from lecture, Sipser, prolem sets, or section, s long s you quote it clerly. The lphet is

More information

CS 301. Lecture 04 Regular Expressions. Stephen Checkoway. January 29, 2018

CS 301. Lecture 04 Regular Expressions. Stephen Checkoway. January 29, 2018 CS 301 Lecture 04 Regulr Expressions Stephen Checkowy Jnury 29, 2018 1 / 35 Review from lst time NFA N = (Q, Σ, δ, q 0, F ) where δ Q Σ P (Q) mps stte nd n lphet symol (or ) to set of sttes We run n NFA

More information

p-adic Egyptian Fractions

p-adic Egyptian Fractions p-adic Egyptin Frctions Contents 1 Introduction 1 2 Trditionl Egyptin Frctions nd Greedy Algorithm 2 3 Set-up 3 4 p-greedy Algorithm 5 5 p-egyptin Trditionl 10 6 Conclusion 1 Introduction An Egyptin frction

More information

12.1 Nondeterminism Nondeterministic Finite Automata. a a b ε. CS125 Lecture 12 Fall 2016

12.1 Nondeterminism Nondeterministic Finite Automata. a a b ε. CS125 Lecture 12 Fall 2016 CS125 Lecture 12 Fll 2016 12.1 Nondeterminism The ide of nondeterministic computtions is to llow our lgorithms to mke guesses, nd only require tht they ccept when the guesses re correct. For exmple, simple

More information

Regular Expressions (RE) Regular Expressions (RE) Regular Expressions (RE) Regular Expressions (RE) Kleene-*

Regular Expressions (RE) Regular Expressions (RE) Regular Expressions (RE) Regular Expressions (RE) Kleene-* Regulr Expressions (RE) Regulr Expressions (RE) Empty set F A RE denotes the empty set Opertion Nottion Lnguge UNIX Empty string A RE denotes the set {} Alterntion R +r L(r ) L(r ) r r Symol Alterntion

More information

CM10196 Topic 4: Functions and Relations

CM10196 Topic 4: Functions and Relations CM096 Topic 4: Functions nd Reltions Guy McCusker W. Functions nd reltions Perhps the most widely used notion in ll of mthemtics is tht of function. Informlly, function is n opertion which tkes n input

More information

CS 310 (sec 20) - Winter Final Exam (solutions) SOLUTIONS

CS 310 (sec 20) - Winter Final Exam (solutions) SOLUTIONS CS 310 (sec 20) - Winter 2003 - Finl Exm (solutions) SOLUTIONS 1. (Logic) Use truth tles to prove the following logicl equivlences: () p q (p p) (q q) () p q (p q) (p q) () p q p q p p q q (q q) (p p)

More information

Deterministic Finite Automata

Deterministic Finite Automata Finite Automt Deterministic Finite Automt H. Geuvers nd J. Rot Institute for Computing nd Informtion Sciences Version: fll 2016 J. Rot Version: fll 2016 Tlen en Automten 1 / 21 Outline Finite Automt Finite

More information

NFAs continued, Closure Properties of Regular Languages

NFAs continued, Closure Properties of Regular Languages Algorithms & Models of Computtion CS/ECE 374, Fll 2017 NFAs continued, Closure Properties of Regulr Lnguges Lecture 5 Tuesdy, Septemer 12, 2017 Sriel Hr-Peled (UIUC) CS374 1 Fll 2017 1 / 31 Regulr Lnguges,

More information

Myhill-Nerode Theorem

Myhill-Nerode Theorem Overview Myhill-Nerode Theorem Correspondence etween DA s nd MN reltions Cnonicl DA for L Computing cnonicl DFA Myhill-Nerode Theorem Deepk D Souz Deprtment of Computer Science nd Automtion Indin Institute

More information

CHAPTER 1 Regular Languages. Contents

CHAPTER 1 Regular Languages. Contents Finite Automt (FA or DFA) CHAPTE 1 egulr Lnguges Contents definitions, exmples, designing, regulr opertions Non-deterministic Finite Automt (NFA) definitions, euivlence of NFAs nd DFAs, closure under regulr

More information

Talen en Automaten Test 1, Mon 7 th Dec, h45 17h30

Talen en Automaten Test 1, Mon 7 th Dec, h45 17h30 Tlen en Automten Test 1, Mon 7 th Dec, 2015 15h45 17h30 This test consists of four exercises over 5 pges. Explin your pproch, nd write your nswer to ech exercise on seprte pge. You cn score mximum of 100

More information

First Midterm Examination

First Midterm Examination Çnky University Deprtment of Computer Engineering 203-204 Fll Semester First Midterm Exmintion ) Design DFA for ll strings over the lphet Σ = {,, c} in which there is no, no nd no cc. 2) Wht lnguge does

More information

First Midterm Examination

First Midterm Examination 24-25 Fll Semester First Midterm Exmintion ) Give the stte digrm of DFA tht recognizes the lnguge A over lphet Σ = {, } where A = {w w contins or } 2) The following DFA recognizes the lnguge B over lphet

More information

Finite-State Automata: Recap

Finite-State Automata: Recap Finite-Stte Automt: Recp Deepk D Souz Deprtment of Computer Science nd Automtion Indin Institute of Science, Bnglore. 09 August 2016 Outline 1 Introduction 2 Forml Definitions nd Nottion 3 Closure under

More information

Grammar. Languages. Content 5/10/16. Automata and Languages. Regular Languages. Regular Languages

Grammar. Languages. Content 5/10/16. Automata and Languages. Regular Languages. Regular Languages 5//6 Grmmr Automt nd Lnguges Regulr Grmmr Context-free Grmmr Context-sensitive Grmmr Prof. Mohmed Hmd Softwre Engineering L. The University of Aizu Jpn Regulr Lnguges Context Free Lnguges Context Sensitive

More information

CS 330 Formal Methods and Models

CS 330 Formal Methods and Models CS 330 Forml Methods nd Models Dn Richrds, George Mson University, Spring 2017 Quiz Solutions Quiz 1, Propositionl Logic Dte: Ferury 2 1. Prove ((( p q) q) p) is tutology () (3pts) y truth tle. p q p q

More information

DFA minimisation using the Myhill-Nerode theorem

DFA minimisation using the Myhill-Nerode theorem DFA minimistion using the Myhill-Nerode theorem Johnn Högerg Lrs Lrsson Astrct The Myhill-Nerode theorem is n importnt chrcteristion of regulr lnguges, nd it lso hs mny prcticl implictions. In this chpter,

More information

Homework 3 Solutions

Homework 3 Solutions CS 341: Foundtions of Computer Science II Prof. Mrvin Nkym Homework 3 Solutions 1. Give NFAs with the specified numer of sttes recognizing ech of the following lnguges. In ll cses, the lphet is Σ = {,1}.

More information

Finite Automata. Informatics 2A: Lecture 3. John Longley. 22 September School of Informatics University of Edinburgh

Finite Automata. Informatics 2A: Lecture 3. John Longley. 22 September School of Informatics University of Edinburgh Lnguges nd Automt Finite Automt Informtics 2A: Lecture 3 John Longley School of Informtics University of Edinburgh jrl@inf.ed.c.uk 22 September 2017 1 / 30 Lnguges nd Automt 1 Lnguges nd Automt Wht is

More information

NFA DFA Example 3 CMSC 330: Organization of Programming Languages. Equivalence of DFAs and NFAs. Equivalence of DFAs and NFAs (cont.

NFA DFA Example 3 CMSC 330: Organization of Programming Languages. Equivalence of DFAs and NFAs. Equivalence of DFAs and NFAs (cont. NFA DFA Exmple 3 CMSC 330: Orgniztion of Progrmming Lnguges NFA {B,D,E {A,E {C,D {E Finite Automt, con't. R = { {A,E, {B,D,E, {C,D, {E 2 Equivlence of DFAs nd NFAs Any string from {A to either {D or {CD

More information

NFAs and Regular Expressions. NFA-ε, continued. Recall. Last class: Today: Fun:

NFAs and Regular Expressions. NFA-ε, continued. Recall. Last class: Today: Fun: CMPU 240 Lnguge Theory nd Computtion Spring 2019 NFAs nd Regulr Expressions Lst clss: Introduced nondeterministic finite utomt with -trnsitions Tody: Prove n NFA- is no more powerful thn n NFA Introduce

More information

12.1 Nondeterminism Nondeterministic Finite Automata. a a b ε. CS125 Lecture 12 Fall 2014

12.1 Nondeterminism Nondeterministic Finite Automata. a a b ε. CS125 Lecture 12 Fall 2014 CS125 Lecture 12 Fll 2014 12.1 Nondeterminism The ide of nondeterministic computtions is to llow our lgorithms to mke guesses, nd only require tht they ccept when the guesses re correct. For exmple, simple

More information

CISC 4090 Theory of Computation

CISC 4090 Theory of Computation 9/6/28 Stereotypicl computer CISC 49 Theory of Computtion Finite stte mchines & Regulr lnguges Professor Dniel Leeds dleeds@fordhm.edu JMH 332 Centrl processing unit (CPU) performs ll the instructions

More information

Scanner. Specifying patterns. Specifying patterns. Operations on languages. A scanner must recognize the units of syntax Some parts are easy:

Scanner. Specifying patterns. Specifying patterns. Operations on languages. A scanner must recognize the units of syntax Some parts are easy: Scnner Specifying ptterns source code tokens scnner prser IR A scnner must recognize the units of syntx Some prts re esy: errors mps chrcters into tokens the sic unit of syntx x = x + y; ecomes

More information

CS 275 Automata and Formal Language Theory

CS 275 Automata and Formal Language Theory CS 275 utomt nd Forml Lnguge Theory Course Notes Prt II: The Recognition Prolem (II) Chpter II.5.: Properties of Context Free Grmmrs (14) nton Setzer (Bsed on ook drft y J. V. Tucker nd K. Stephenson)

More information

1. For each of the following theorems, give a two or three sentence sketch of how the proof goes or why it is not true.

1. For each of the following theorems, give a two or three sentence sketch of how the proof goes or why it is not true. York University CSE 2 Unit 3. DFA Clsses Converting etween DFA, NFA, Regulr Expressions, nd Extended Regulr Expressions Instructor: Jeff Edmonds Don t chet y looking t these nswers premturely.. For ech

More information

Let's start with an example:

Let's start with an example: Finite Automt Let's strt with n exmple: Here you see leled circles tht re sttes, nd leled rrows tht re trnsitions. One of the sttes is mrked "strt". One of the sttes hs doule circle; this is terminl stte

More information

CS 311 Homework 3 due 16:30, Thursday, 14 th October 2010

CS 311 Homework 3 due 16:30, Thursday, 14 th October 2010 CS 311 Homework 3 due 16:30, Thursdy, 14 th Octoer 2010 Homework must e sumitted on pper, in clss. Question 1. [15 pts.; 5 pts. ech] Drw stte digrms for NFAs recognizing the following lnguges:. L = {w

More information

CMSC 330: Organization of Programming Languages. DFAs, and NFAs, and Regexps (Oh my!)

CMSC 330: Organization of Programming Languages. DFAs, and NFAs, and Regexps (Oh my!) CMSC 330: Orgniztion of Progrmming Lnguges DFAs, nd NFAs, nd Regexps (Oh my!) CMSC330 Spring 2018 Types of Finite Automt Deterministic Finite Automt (DFA) Exctly one sequence of steps for ech string All

More information

7.2 The Definite Integral

7.2 The Definite Integral 7.2 The Definite Integrl the definite integrl In the previous section, it ws found tht if function f is continuous nd nonnegtive, then the re under the grph of f on [, b] is given by F (b) F (), where

More information

Theory of Computation Regular Languages. (NTU EE) Regular Languages Fall / 38

Theory of Computation Regular Languages. (NTU EE) Regular Languages Fall / 38 Theory of Computtion Regulr Lnguges (NTU EE) Regulr Lnguges Fll 2017 1 / 38 Schemtic of Finite Automt control 0 0 1 0 1 1 1 0 Figure: Schemtic of Finite Automt A finite utomton hs finite set of control

More information

Anatomy of a Deterministic Finite Automaton. Deterministic Finite Automata. A machine so simple that you can understand it in less than one minute

Anatomy of a Deterministic Finite Automaton. Deterministic Finite Automata. A machine so simple that you can understand it in less than one minute Victor Admchik Dnny Sletor Gret Theoreticl Ides In Computer Science CS 5-25 Spring 2 Lecture 2 Mr 3, 2 Crnegie Mellon University Deterministic Finite Automt Finite Automt A mchine so simple tht you cn

More information

A negative answer to a question of Wilke on varieties of!-languages

A negative answer to a question of Wilke on varieties of!-languages A negtive nswer to question of Wilke on vrieties of!-lnguges Jen-Eric Pin () Astrct. In recent pper, Wilke sked whether the oolen comintions of!-lnguges of the form! L, for L in given +-vriety of lnguges,

More information

CS 373, Spring Solutions to Mock midterm 1 (Based on first midterm in CS 273, Fall 2008.)

CS 373, Spring Solutions to Mock midterm 1 (Based on first midterm in CS 273, Fall 2008.) CS 373, Spring 29. Solutions to Mock midterm (sed on first midterm in CS 273, Fll 28.) Prolem : Short nswer (8 points) The nswers to these prolems should e short nd not complicted. () If n NF M ccepts

More information

1B40 Practical Skills

1B40 Practical Skills B40 Prcticl Skills Comining uncertinties from severl quntities error propgtion We usully encounter situtions where the result of n experiment is given in terms of two (or more) quntities. We then need

More information

The Regulated and Riemann Integrals

The Regulated and Riemann Integrals Chpter 1 The Regulted nd Riemnn Integrls 1.1 Introduction We will consider severl different pproches to defining the definite integrl f(x) dx of function f(x). These definitions will ll ssign the sme vlue

More information

Lexical Analysis Finite Automate

Lexical Analysis Finite Automate Lexicl Anlysis Finite Automte CMPSC 470 Lecture 04 Topics: Deterministic Finite Automt (DFA) Nondeterministic Finite Automt (NFA) Regulr Expression NFA DFA A. Finite Automt (FA) FA re grph, like trnsition

More information

A compact proof of decidability for regular expression equivalence

A compact proof of decidability for regular expression equivalence A compct proof of decidbility for regulr expression equivlence ITP 2012 Princeton, USA Andre Asperti Deprtment of Computer Science University of Bologn 25/08/2011 Abstrct We introduce the notion of pointed

More information

Quadratic Forms. Quadratic Forms

Quadratic Forms. Quadratic Forms Qudrtic Forms Recll the Simon & Blume excerpt from n erlier lecture which sid tht the min tsk of clculus is to pproximte nonliner functions with liner functions. It s ctully more ccurte to sy tht we pproximte

More information

Greedy regular expression matching

Greedy regular expression matching Alin Frisch INRIA Luc Crdelli MSRC 2004-05-15 ICALP The mtching prolem The prolem Project the structure of regulr expression on flt sequence. R = ( ) w = 1 2 1 2 3 v = [1 : [ 1 ; 2 ]; 2 : 1 ; 2 : 2 ; 1

More information

NFAs continued, Closure Properties of Regular Languages

NFAs continued, Closure Properties of Regular Languages lgorithms & Models of omputtion S/EE 374, Spring 209 NFs continued, losure Properties of Regulr Lnguges Lecture 5 Tuesdy, Jnury 29, 209 Regulr Lnguges, DFs, NFs Lnguges ccepted y DFs, NFs, nd regulr expressions

More information

Theory of Computation Regular Languages

Theory of Computation Regular Languages Theory of Computtion Regulr Lnguges Bow-Yw Wng Acdemi Sinic Spring 2012 Bow-Yw Wng (Acdemi Sinic) Regulr Lnguges Spring 2012 1 / 38 Schemtic of Finite Automt control 0 0 1 0 1 1 1 0 Figure: Schemtic of

More information

Type Theory. Coinduction in Type Theory. Andreas Abel. Department of Computer Science and Engineering Chalmers and Gothenburg University

Type Theory. Coinduction in Type Theory. Andreas Abel. Department of Computer Science and Engineering Chalmers and Gothenburg University Type Theory Coinduction in Type Theory Andres Ael Deprtment of Computer Science nd Engineering Chlmers nd Gothenurg University Type Theory Course CM0859 (2017-1) Universidd EAFIT, Medellin, Colomi 6-10

More information

1 From NFA to regular expression

1 From NFA to regular expression Note 1: How to convert DFA/NFA to regulr expression Version: 1.0 S/EE 374, Fll 2017 Septemer 11, 2017 In this note, we show tht ny DFA cn e converted into regulr expression. Our construction would work

More information

Homework Solution - Set 5 Due: Friday 10/03/08

Homework Solution - Set 5 Due: Friday 10/03/08 CE 96 Introduction to the Theory of Computtion ll 2008 Homework olution - et 5 Due: ridy 10/0/08 1. Textook, Pge 86, Exercise 1.21. () 1 2 Add new strt stte nd finl stte. Mke originl finl stte non-finl.

More information

Improper Integrals. The First Fundamental Theorem of Calculus, as we ve discussed in class, goes as follows:

Improper Integrals. The First Fundamental Theorem of Calculus, as we ve discussed in class, goes as follows: Improper Integrls The First Fundmentl Theorem of Clculus, s we ve discussed in clss, goes s follows: If f is continuous on the intervl [, ] nd F is function for which F t = ft, then ftdt = F F. An integrl

More information

Genetic Programming. Outline. Evolutionary Strategies. Evolutionary strategies Genetic programming Summary

Genetic Programming. Outline. Evolutionary Strategies. Evolutionary strategies Genetic programming Summary Outline Genetic Progrmming Evolutionry strtegies Genetic progrmming Summry Bsed on the mteril provided y Professor Michel Negnevitsky Evolutionry Strtegies An pproch simulting nturl evolution ws proposed

More information

Boolean algebra.

Boolean algebra. http://en.wikipedi.org/wiki/elementry_boolen_lger Boolen lger www.tudorgir.com Computer science is not out computers, it is out computtion nd informtion. computtion informtion computer informtion Turing

More information

The size of subsequence automaton

The size of subsequence automaton Theoreticl Computer Science 4 (005) 79 84 www.elsevier.com/locte/tcs Note The size of susequence utomton Zdeněk Troníček,, Ayumi Shinohr,c Deprtment of Computer Science nd Engineering, FEE CTU in Prgue,

More information

Closure Properties of Regular Languages

Closure Properties of Regular Languages Closure Properties of Regulr Lnguges Regulr lnguges re closed under mny set opertions. Let L 1 nd L 2 e regulr lnguges. (1) L 1 L 2 (the union) is regulr. (2) L 1 L 2 (the conctention) is regulr. (3) L

More information

Math 1B, lecture 4: Error bounds for numerical methods

Math 1B, lecture 4: Error bounds for numerical methods Mth B, lecture 4: Error bounds for numericl methods Nthn Pflueger 4 September 0 Introduction The five numericl methods descried in the previous lecture ll operte by the sme principle: they pproximte the

More information

and that at t = 0 the object is at position 5. Find the position of the object at t = 2.

and that at t = 0 the object is at position 5. Find the position of the object at t = 2. 7.2 The Fundmentl Theorem of Clculus 49 re mny, mny problems tht pper much different on the surfce but tht turn out to be the sme s these problems, in the sense tht when we try to pproimte solutions we

More information

Homework 4. 0 ε 0. (00) ε 0 ε 0 (00) (11) CS 341: Foundations of Computer Science II Prof. Marvin Nakayama

Homework 4. 0 ε 0. (00) ε 0 ε 0 (00) (11) CS 341: Foundations of Computer Science II Prof. Marvin Nakayama CS 341: Foundtions of Computer Science II Prof. Mrvin Nkym Homework 4 1. UsetheproceduredescriedinLemm1.55toconverttheregulrexpression(((00) (11)) 01) into n NFA. Answer: 0 0 1 1 00 0 0 11 1 1 01 0 1 (00)

More information

CS375: Logic and Theory of Computing

CS375: Logic and Theory of Computing CS375: Logic nd Theory of Computing Fuhu (Frnk) Cheng Deprtment of Computer Science University of Kentucky 1 Tle of Contents: Week 1: Preliminries (set lger, reltions, functions) (red Chpters 1-4) Weeks

More information

CS 267: Automated Verification. Lecture 8: Automata Theoretic Model Checking. Instructor: Tevfik Bultan

CS 267: Automated Verification. Lecture 8: Automata Theoretic Model Checking. Instructor: Tevfik Bultan CS 267: Automted Verifiction Lecture 8: Automt Theoretic Model Checking Instructor: Tevfik Bultn LTL Properties Büchi utomt [Vrdi nd Wolper LICS 86] Büchi utomt: Finite stte utomt tht ccept infinite strings

More information

CSCI 340: Computational Models. Transition Graphs. Department of Computer Science

CSCI 340: Computational Models. Transition Graphs. Department of Computer Science CSCI 340: Computtionl Models Trnsition Grphs Chpter 6 Deprtment of Computer Science Relxing Restrints on Inputs We cn uild n FA tht ccepts only the word! 5 sttes ecuse n FA cn only process one letter t

More information