author | Christian Urban <urbanc@in.tum.de> |
Tue, 20 Aug 2019 23:42:28 +0200 | |
changeset 343 | f139bdc0dcd5 |
parent 305 | 6e2cef17a9b3 |
permissions | -rw-r--r-- |
298 | 1 |
|
2 |
import scala.language.implicitConversions |
|
3 |
import scala.language.reflectiveCalls |
|
4 |
import scala.annotation.tailrec |
|
5 |
import scala.util.Try |
|
6 |
||
7 |
def escape(raw: String) : String = { |
|
8 |
import scala.reflect.runtime.universe._ |
|
9 |
Literal(Constant(raw)).toString |
|
10 |
} |
|
11 |
||
12 |
def esc2(r: (String, String)) = (escape(r._1), escape(r._2)) |
|
13 |
||
14 |
def distinctBy[B, C](xs: List[B], f: B => C, acc: List[C] = Nil): List[B] = xs match { |
|
15 |
case Nil => Nil |
|
16 |
case (x::xs) => { |
|
17 |
val res = f(x) |
|
18 |
if (acc.contains(res)) distinctBy(xs, f, acc) |
|
19 |
else x::distinctBy(xs, f, res::acc) |
|
20 |
} |
|
21 |
} |
|
22 |
||
23 |
abstract class Bit |
|
24 |
case object Z extends Bit |
|
25 |
case object S extends Bit |
|
26 |
case class C(c: Char) extends Bit |
|
27 |
||
28 |
type Bits = List[Bit] |
|
29 |
||
30 |
// usual regular expressions |
|
31 |
abstract class Rexp |
|
32 |
case object ZERO extends Rexp |
|
33 |
case object ONE extends Rexp |
|
305 | 34 |
case class PRED(f: Char => Boolean, s: String = "_") extends Rexp |
298 | 35 |
case class ALTS(rs: List[Rexp]) extends Rexp |
36 |
case class SEQ(r1: Rexp, r2: Rexp) extends Rexp |
|
37 |
case class STAR(r: Rexp) extends Rexp |
|
38 |
case class RECD(x: String, r: Rexp) extends Rexp |
|
39 |
||
40 |
||
41 |
// abbreviations |
|
305 | 42 |
def CHAR(c: Char) = PRED(_ == c, c.toString) |
298 | 43 |
def ALT(r1: Rexp, r2: Rexp) = ALTS(List(r1, r2)) |
44 |
def PLUS(r: Rexp) = SEQ(r, STAR(r)) |
|
45 |
||
46 |
// annotated regular expressions |
|
47 |
abstract class ARexp |
|
48 |
case object AZERO extends ARexp |
|
49 |
case class AONE(bs: Bits) extends ARexp |
|
305 | 50 |
case class APRED(bs: Bits, f: Char => Boolean, s: String = "_") extends ARexp |
298 | 51 |
case class AALTS(bs: Bits, rs: List[ARexp]) extends ARexp |
52 |
case class ASEQ(bs: Bits, r1: ARexp, r2: ARexp) extends ARexp |
|
53 |
case class ASTAR(bs: Bits, r: ARexp) extends ARexp |
|
54 |
||
55 |
// abbreviations |
|
56 |
def AALT(bs: Bits, r1: ARexp, r2: ARexp) = AALTS(bs, List(r1, r2)) |
|
57 |
||
58 |
// values |
|
59 |
abstract class Val |
|
60 |
case object Empty extends Val |
|
61 |
case class Chr(c: Char) extends Val |
|
62 |
case class Sequ(v1: Val, v2: Val) extends Val |
|
63 |
case class Left(v: Val) extends Val |
|
64 |
case class Right(v: Val) extends Val |
|
65 |
case class Stars(vs: List[Val]) extends Val |
|
66 |
case class Rec(x: String, v: Val) extends Val |
|
67 |
||
68 |
||
69 |
||
70 |
// some convenience for typing in regular expressions |
|
71 |
def charlist2rexp(s : List[Char]): Rexp = s match { |
|
72 |
case Nil => ONE |
|
73 |
case c::Nil => CHAR(c) |
|
74 |
case c::s => SEQ(CHAR(c), charlist2rexp(s)) |
|
75 |
} |
|
76 |
implicit def string2rexp(s : String) : Rexp = charlist2rexp(s.toList) |
|
77 |
||
78 |
implicit def RexpOps(r: Rexp) = new { |
|
79 |
def | (s: Rexp) = ALT(r, s) |
|
80 |
def % = STAR(r) |
|
81 |
def ~ (s: Rexp) = SEQ(r, s) |
|
82 |
} |
|
83 |
||
84 |
implicit def stringOps(s: String) = new { |
|
85 |
def | (r: Rexp) = ALT(s, r) |
|
86 |
def | (r: String) = ALT(s, r) |
|
87 |
def % = STAR(s) |
|
88 |
def ~ (r: Rexp) = SEQ(s, r) |
|
89 |
def ~ (r: String) = SEQ(s, r) |
|
90 |
def $ (r: Rexp) = RECD(s, r) |
|
91 |
} |
|
92 |
||
93 |
||
300 | 94 |
// string of a regular expressions - for testing purposes |
305 | 95 |
def string(r: Rexp): String = r match { |
96 |
case ZERO => "0" |
|
97 |
case ONE => "1" |
|
98 |
case PRED(_, s) => s |
|
99 |
case ALTS(rs) => rs.map(string).mkString("[", "|", "]") |
|
100 |
case SEQ(r1, r2) => s"(${string(r1)} ~ ${string(r2)})" |
|
101 |
case STAR(r) => s"{${string(r)}}*" |
|
102 |
case RECD(x, r) => s"(${x}! ${string(r)})" |
|
103 |
} |
|
104 |
||
105 |
// string of an annotated regular expressions - for testing purposes |
|
106 |
||
107 |
def astring(a: ARexp): String = a match { |
|
108 |
case AZERO => "0" |
|
109 |
case AONE(_) => "1" |
|
110 |
case APRED(_, _, s) => s |
|
111 |
case AALTS(_, rs) => rs.map(astring).mkString("[", "|", "]") |
|
112 |
case ASEQ(_, r1, r2) => s"(${astring(r1)} ~ ${astring(r2)})" |
|
113 |
case ASTAR(_, r) => s"{${astring(r)}}*" |
|
114 |
} |
|
115 |
||
300 | 116 |
|
298 | 117 |
//-------------------------------------------------------------------------------------------------------- |
118 |
// START OF NON-BITCODE PART |
|
119 |
// |
|
120 |
||
121 |
// nullable function: tests whether the regular |
|
122 |
// expression can recognise the empty string |
|
123 |
def nullable (r: Rexp) : Boolean = r match { |
|
124 |
case ZERO => false |
|
125 |
case ONE => true |
|
305 | 126 |
case PRED(_, _) => false |
298 | 127 |
case ALTS(rs) => rs.exists(nullable) |
128 |
case SEQ(r1, r2) => nullable(r1) && nullable(r2) |
|
129 |
case STAR(_) => true |
|
130 |
case RECD(_, r) => nullable(r) |
|
131 |
} |
|
132 |
||
133 |
// derivative of a regular expression w.r.t. a character |
|
134 |
def der (c: Char, r: Rexp) : Rexp = r match { |
|
135 |
case ZERO => ZERO |
|
136 |
case ONE => ZERO |
|
305 | 137 |
case PRED(f, _) => if (f(c)) ONE else ZERO |
298 | 138 |
case ALTS(List(r1, r2)) => ALTS(List(der(c, r1), der(c, r2))) |
139 |
case SEQ(r1, r2) => |
|
140 |
if (nullable(r1)) ALTS(List(SEQ(der(c, r1), r2), der(c, r2))) |
|
141 |
else SEQ(der(c, r1), r2) |
|
142 |
case STAR(r) => SEQ(der(c, r), STAR(r)) |
|
143 |
case RECD(_, r1) => der(c, r1) |
|
144 |
} |
|
145 |
||
146 |
||
147 |
def flatten(v: Val) : String = v match { |
|
148 |
case Empty => "" |
|
149 |
case Chr(c) => c.toString |
|
150 |
case Left(v) => flatten(v) |
|
151 |
case Right(v) => flatten(v) |
|
152 |
case Sequ(v1, v2) => flatten(v1) + flatten(v2) |
|
153 |
case Stars(vs) => vs.map(flatten).mkString |
|
154 |
case Rec(_, v) => flatten(v) |
|
155 |
} |
|
156 |
||
157 |
// extracts an environment from a value |
|
158 |
def env(v: Val) : List[(String, String)] = v match { |
|
159 |
case Empty => Nil |
|
160 |
case Chr(c) => Nil |
|
161 |
case Left(v) => env(v) |
|
162 |
case Right(v) => env(v) |
|
163 |
case Sequ(v1, v2) => env(v1) ::: env(v2) |
|
164 |
case Stars(vs) => vs.flatMap(env) |
|
165 |
case Rec(x, v) => (x, flatten(v))::env(v) |
|
166 |
} |
|
167 |
||
168 |
||
169 |
// injection part |
|
170 |
def mkeps(r: Rexp) : Val = r match { |
|
171 |
case ONE => Empty |
|
172 |
case ALTS(List(r1, r2)) => |
|
173 |
if (nullable(r1)) Left(mkeps(r1)) else Right(mkeps(r2)) |
|
174 |
case SEQ(r1, r2) => Sequ(mkeps(r1), mkeps(r2)) |
|
175 |
case STAR(r) => Stars(Nil) |
|
176 |
case RECD(x, r) => Rec(x, mkeps(r)) |
|
177 |
} |
|
178 |
||
179 |
def inj(r: Rexp, c: Char, v: Val) : Val = (r, v) match { |
|
180 |
case (STAR(r), Sequ(v1, Stars(vs))) => Stars(inj(r, c, v1)::vs) |
|
181 |
case (SEQ(r1, r2), Sequ(v1, v2)) => Sequ(inj(r1, c, v1), v2) |
|
182 |
case (SEQ(r1, r2), Left(Sequ(v1, v2))) => Sequ(inj(r1, c, v1), v2) |
|
183 |
case (SEQ(r1, r2), Right(v2)) => Sequ(mkeps(r1), inj(r2, c, v2)) |
|
184 |
case (ALTS(List(r1, r2)), Left(v1)) => Left(inj(r1, c, v1)) |
|
185 |
case (ALTS(List(r1, r2)), Right(v2)) => Right(inj(r2, c, v2)) |
|
305 | 186 |
case (PRED(_, _), Empty) => Chr(c) |
298 | 187 |
case (RECD(x, r1), _) => Rec(x, inj(r1, c, v)) |
188 |
} |
|
189 |
||
190 |
// lexing without simplification |
|
191 |
def lex(r: Rexp, s: List[Char]) : Val = s match { |
|
192 |
case Nil => if (nullable(r)) mkeps(r) else throw new Exception("Not matched") |
|
193 |
case c::cs => inj(r, c, lex(der(c, r), cs)) |
|
194 |
} |
|
195 |
||
196 |
def lexing(r: Rexp, s: String) : Val = lex(r, s.toList) |
|
197 |
||
198 |
//println(lexing(("ab" | "ab") ~ ("b" | ONE), "ab")) |
|
199 |
||
200 |
// some "rectification" functions for simplification |
|
201 |
def F_ID(v: Val): Val = v |
|
202 |
def F_RIGHT(f: Val => Val) = (v:Val) => Right(f(v)) |
|
203 |
def F_LEFT(f: Val => Val) = (v:Val) => Left(f(v)) |
|
204 |
def F_ALT(f1: Val => Val, f2: Val => Val) = (v:Val) => v match { |
|
205 |
case Right(v) => Right(f2(v)) |
|
206 |
case Left(v) => Left(f1(v)) |
|
207 |
} |
|
208 |
def F_SEQ(f1: Val => Val, f2: Val => Val) = (v:Val) => v match { |
|
209 |
case Sequ(v1, v2) => Sequ(f1(v1), f2(v2)) |
|
210 |
} |
|
211 |
def F_SEQ_Empty1(f1: Val => Val, f2: Val => Val) = |
|
212 |
(v:Val) => Sequ(f1(Empty), f2(v)) |
|
213 |
def F_SEQ_Empty2(f1: Val => Val, f2: Val => Val) = |
|
214 |
(v:Val) => Sequ(f1(v), f2(Empty)) |
|
215 |
def F_RECD(f: Val => Val) = (v:Val) => v match { |
|
216 |
case Rec(x, v) => Rec(x, f(v)) |
|
217 |
} |
|
218 |
def F_ERROR(v: Val): Val = throw new Exception("error") |
|
219 |
||
220 |
// simplification of regular expressions returning also an |
|
221 |
// rectification function; no simplification under STAR |
|
305 | 222 |
|
223 |
var profile_simp : Int = 0 |
|
224 |
||
225 |
def simp(r: Rexp): (Rexp, Val => Val) = { |
|
226 |
profile_simp = profile_simp + 1 |
|
227 |
r match { |
|
298 | 228 |
case ALTS(List(r1, r2)) => { |
229 |
val (r1s, f1s) = simp(r1) |
|
230 |
val (r2s, f2s) = simp(r2) |
|
231 |
(r1s, r2s) match { |
|
232 |
case (ZERO, _) => (r2s, F_RIGHT(f2s)) |
|
233 |
case (_, ZERO) => (r1s, F_LEFT(f1s)) |
|
234 |
case _ => if (r1s == r2s) (r1s, F_LEFT(f1s)) |
|
235 |
else (ALTS(List(r1s, r2s)), F_ALT(f1s, f2s)) |
|
236 |
} |
|
237 |
} |
|
238 |
case SEQ(r1, r2) => { |
|
239 |
val (r1s, f1s) = simp(r1) |
|
240 |
val (r2s, f2s) = simp(r2) |
|
241 |
(r1s, r2s) match { |
|
242 |
case (ZERO, _) => (ZERO, F_ERROR) |
|
243 |
case (_, ZERO) => (ZERO, F_ERROR) |
|
244 |
case (ONE, _) => (r2s, F_SEQ_Empty1(f1s, f2s)) |
|
245 |
case (_, ONE) => (r1s, F_SEQ_Empty2(f1s, f2s)) |
|
246 |
case _ => (SEQ(r1s,r2s), F_SEQ(f1s, f2s)) |
|
247 |
} |
|
248 |
} |
|
249 |
case RECD(x, r1) => { |
|
250 |
val (r1s, f1s) = simp(r1) |
|
251 |
(RECD(x, r1s), F_RECD(f1s)) |
|
252 |
} |
|
253 |
case r => (r, F_ID) |
|
305 | 254 |
}} |
298 | 255 |
|
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
256 |
def ders_simp(s: List[Char], r: Rexp) : Rexp = s match { |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
257 |
case Nil => r |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
258 |
case c::s => ders_simp(s, simp(der(c, r))._1) |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
259 |
} |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
260 |
|
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
261 |
|
298 | 262 |
def lex_simp(r: Rexp, s: List[Char]) : Val = s match { |
263 |
case Nil => if (nullable(r)) mkeps(r) else throw new Exception("Not matched") |
|
264 |
case c::cs => { |
|
265 |
val (r_simp, f_simp) = simp(der(c, r)) |
|
266 |
inj(r, c, f_simp(lex_simp(r_simp, cs))) |
|
267 |
} |
|
268 |
} |
|
269 |
||
270 |
def lexing_simp(r: Rexp, s: String) : Val = lex_simp(r, s.toList) |
|
271 |
||
272 |
println(lexing_simp(("a" | "ab") ~ ("b" | ""), "ab")) |
|
273 |
||
274 |
||
275 |
def tokenise_simp(r: Rexp, s: String) = env(lexing_simp(r, s)).map(esc2) |
|
276 |
||
277 |
//-------------------------------------------------------------------------------------------------------- |
|
278 |
// BITCODED PART |
|
279 |
||
280 |
||
281 |
def fuse(bs: Bits, r: ARexp) : ARexp = r match { |
|
282 |
case AZERO => AZERO |
|
283 |
case AONE(cs) => AONE(bs ++ cs) |
|
305 | 284 |
case APRED(cs, f, s) => APRED(bs ++ cs, f, s) |
298 | 285 |
case AALTS(cs, rs) => AALTS(bs ++ cs, rs) |
286 |
case ASEQ(cs, r1, r2) => ASEQ(bs ++ cs, r1, r2) |
|
287 |
case ASTAR(cs, r) => ASTAR(bs ++ cs, r) |
|
288 |
} |
|
289 |
||
290 |
// translation into ARexps |
|
291 |
def internalise(r: Rexp) : ARexp = r match { |
|
292 |
case ZERO => AZERO |
|
293 |
case ONE => AONE(Nil) |
|
305 | 294 |
case PRED(f, s) => APRED(Nil, f, s) |
298 | 295 |
case ALTS(List(r1, r2)) => |
296 |
AALTS(Nil, List(fuse(List(Z), internalise(r1)), fuse(List(S), internalise(r2)))) |
|
297 |
case ALTS(r1::rs) => { |
|
298 |
val AALTS(Nil, rs2) = internalise(ALTS(rs)) |
|
299 |
AALTS(Nil, fuse(List(Z), internalise(r1)) :: rs2.map(fuse(List(S), _))) |
|
300 |
} |
|
301 |
case SEQ(r1, r2) => ASEQ(Nil, internalise(r1), internalise(r2)) |
|
302 |
case STAR(r) => ASTAR(Nil, internalise(r)) |
|
303 |
case RECD(x, r) => internalise(r) |
|
304 |
} |
|
305 |
||
306 |
internalise(("a" | "ab") ~ ("b" | "")) |
|
307 |
||
308 |
// decoding of values from bit sequences |
|
309 |
def decode_aux(r: Rexp, bs: Bits) : (Val, Bits) = (r, bs) match { |
|
310 |
case (ONE, bs) => (Empty, bs) |
|
305 | 311 |
case (PRED(f, _), C(c)::bs) => (Chr(c), bs) |
298 | 312 |
case (ALTS(r::Nil), bs) => decode_aux(r, bs) |
313 |
case (ALTS(rs), bs) => bs match { |
|
314 |
case Z::bs1 => { |
|
315 |
val (v, bs2) = decode_aux(rs.head, bs1) |
|
316 |
(Left(v), bs2) |
|
317 |
} |
|
318 |
case S::bs1 => { |
|
319 |
val (v, bs2) = decode_aux(ALTS(rs.tail), bs1) |
|
320 |
(Right(v), bs2) |
|
321 |
} |
|
322 |
} |
|
323 |
case (SEQ(r1, r2), bs) => { |
|
324 |
val (v1, bs1) = decode_aux(r1, bs) |
|
325 |
val (v2, bs2) = decode_aux(r2, bs1) |
|
326 |
(Sequ(v1, v2), bs2) |
|
327 |
} |
|
328 |
case (STAR(r1), S::bs) => { |
|
329 |
val (v, bs1) = decode_aux(r1, bs) |
|
330 |
val (Stars(vs), bs2) = decode_aux(STAR(r1), bs1) |
|
331 |
(Stars(v::vs), bs2) |
|
332 |
} |
|
333 |
case (STAR(_), Z::bs) => (Stars(Nil), bs) |
|
334 |
case (RECD(x, r1), bs) => { |
|
335 |
val (v, bs1) = decode_aux(r1, bs) |
|
336 |
(Rec(x, v), bs1) |
|
337 |
} |
|
338 |
} |
|
339 |
||
340 |
def decode(r: Rexp, bs: Bits) = decode_aux(r, bs) match { |
|
341 |
case (v, Nil) => v |
|
342 |
case _ => throw new Exception("Not decodable") |
|
343 |
} |
|
344 |
||
345 |
||
346 |
//erase function: extracts a Rexp from Arexp |
|
347 |
def erase(r: ARexp) : Rexp = r match{ |
|
348 |
case AZERO => ZERO |
|
349 |
case AONE(_) => ONE |
|
305 | 350 |
case APRED(bs, f, s) => PRED(f, s) |
298 | 351 |
case AALTS(bs, rs) => ALTS(rs.map(erase(_))) |
352 |
case ASEQ(bs, r1, r2) => SEQ (erase(r1), erase(r2)) |
|
353 |
case ASTAR(cs, r)=> STAR(erase(r)) |
|
354 |
} |
|
355 |
||
356 |
||
357 |
// bnullable function: tests whether the aregular |
|
358 |
// expression can recognise the empty string |
|
359 |
def bnullable (r: ARexp) : Boolean = r match { |
|
360 |
case AZERO => false |
|
361 |
case AONE(_) => true |
|
305 | 362 |
case APRED(_,_,_) => false |
298 | 363 |
case AALTS(_, rs) => rs.exists(bnullable) |
364 |
case ASEQ(_, r1, r2) => bnullable(r1) && bnullable(r2) |
|
365 |
case ASTAR(_, _) => true |
|
366 |
} |
|
367 |
||
368 |
def bmkeps(r: ARexp) : Bits = r match { |
|
369 |
case AONE(bs) => bs |
|
370 |
case AALTS(bs, rs) => { |
|
371 |
val n = rs.indexWhere(bnullable) |
|
372 |
bs ++ bmkeps(rs(n)) |
|
373 |
} |
|
374 |
case ASEQ(bs, r1, r2) => bs ++ bmkeps(r1) ++ bmkeps(r2) |
|
375 |
case ASTAR(bs, r) => bs ++ List(Z) |
|
376 |
} |
|
377 |
||
378 |
// derivative of a regular expression w.r.t. a character |
|
379 |
def bder(c: Char, r: ARexp) : ARexp = r match { |
|
380 |
case AZERO => AZERO |
|
381 |
case AONE(_) => AZERO |
|
305 | 382 |
case APRED(bs, f, _) => if (f(c)) AONE(bs:::List(C(c))) else AZERO |
298 | 383 |
case AALTS(bs, rs) => AALTS(bs, rs.map(bder(c, _))) |
384 |
case ASEQ(bs, r1, r2) => |
|
385 |
if (bnullable(r1)) AALT(bs, ASEQ(Nil, bder(c, r1), r2), fuse(bmkeps(r1), bder(c, r2))) |
|
386 |
else ASEQ(bs, bder(c, r1), r2) |
|
387 |
case ASTAR(bs, r) => ASEQ(bs, fuse(List(S), bder(c, r)), ASTAR(Nil, r)) |
|
388 |
} |
|
389 |
||
390 |
||
391 |
// derivative w.r.t. a string (iterates bder) |
|
392 |
@tailrec |
|
393 |
def bders (s: List[Char], r: ARexp) : ARexp = s match { |
|
394 |
case Nil => r |
|
395 |
case c::s => bders(s, bder(c, r)) |
|
396 |
} |
|
397 |
||
398 |
||
305 | 399 |
def inc_profile(m: Map[ARexp, Int], a: ARexp) : Map[ARexp, Int] = { |
400 |
val current = m.getOrElse(a, 0) |
|
401 |
m + (a -> (current + 1)) |
|
402 |
} |
|
403 |
||
404 |
var profile_bsimp_args : Map[ARexp, Int] = Map() |
|
405 |
var profile_bsimp : Int = 0 |
|
406 |
var profile_flats : Int = 0 |
|
407 |
||
408 |
||
409 |
def flats(rs: List[ARexp]): List[ARexp] = { |
|
410 |
profile_flats = profile_flats + 1 |
|
411 |
rs match { |
|
298 | 412 |
case Nil => Nil |
413 |
case AZERO :: rs1 => flats(rs1) |
|
414 |
case AALTS(bs, rs1) :: rs2 => rs1.map(fuse(bs, _)) ::: flats(rs2) |
|
415 |
case r1 :: rs2 => r1 :: flats(rs2) |
|
305 | 416 |
}} |
298 | 417 |
|
305 | 418 |
def bsimp(r: ARexp): ARexp = { |
419 |
//profile_bsimp = inc_profile(profile_bsimp, r) |
|
420 |
profile_bsimp = profile_bsimp + 1 |
|
421 |
r match { |
|
298 | 422 |
case ASEQ(bs1, r1, r2) => (bsimp(r1), bsimp(r2)) match { |
423 |
case (AZERO, _) => AZERO |
|
424 |
case (_, AZERO) => AZERO |
|
425 |
case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s) |
|
426 |
case (r1s, r2s) => ASEQ(bs1, r1s, r2s) |
|
427 |
} |
|
428 |
case AALTS(bs1, rs) => distinctBy(flats(rs.map(bsimp)), erase) match { |
|
429 |
case Nil => AZERO |
|
300 | 430 |
case r :: Nil => fuse(bs1, r) |
298 | 431 |
case rs => AALTS(bs1, rs) |
432 |
} |
|
300 | 433 |
//case ASTAR(bs1, r1) => ASTAR(bs1, bsimp(r1)) |
298 | 434 |
case r => r |
305 | 435 |
} |
298 | 436 |
} |
437 |
||
300 | 438 |
|
439 |
||
305 | 440 |
|
298 | 441 |
def bders_simp (s: List[Char], r: ARexp) : ARexp = s match { |
442 |
case Nil => r |
|
443 |
case c::s => bders_simp(s, bsimp(bder(c, r))) |
|
444 |
} |
|
445 |
||
446 |
def blex_simp(r: ARexp, s: List[Char]) : Bits = s match { |
|
447 |
case Nil => if (bnullable(r)) bmkeps(r) |
|
448 |
else throw new Exception("Not matched") |
|
449 |
case c::cs => blex_simp(bsimp(bder(c, r)), cs) |
|
450 |
} |
|
451 |
||
452 |
||
453 |
def blexing_simp(r: Rexp, s: String) : Val = |
|
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
454 |
decode(r, blex_simp(internalise(r), s.toList)) |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
455 |
|
298 | 456 |
|
457 |
def btokenise_simp(r: Rexp, s: String) = env(blexing_simp(r, s)).map(esc2) |
|
458 |
||
459 |
||
460 |
||
300 | 461 |
// INCLUDING SIMPLIFICATION UNDER STARS |
462 |
||
463 |
def bsimp_full(r: ARexp): ARexp = r match { |
|
464 |
case ASEQ(bs1, r1, r2) => (bsimp_full(r1), bsimp_full(r2)) match { |
|
465 |
case (AZERO, _) => AZERO |
|
466 |
case (_, AZERO) => AZERO |
|
467 |
case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s) |
|
468 |
case (r1s, r2s) => ASEQ(bs1, r1s, r2s) |
|
469 |
} |
|
470 |
case AALTS(bs1, rs) => distinctBy(flats(rs.map(bsimp_full)), erase) match { |
|
471 |
case Nil => AZERO |
|
472 |
case r :: Nil => fuse(bs1, r) |
|
473 |
case rs => AALTS(bs1, rs) |
|
474 |
} |
|
475 |
case ASTAR(bs1, r1) => ASTAR(bs1, bsimp_full(r1)) |
|
476 |
case r => r |
|
477 |
} |
|
478 |
||
479 |
def bders_simp_full(s: List[Char], r: ARexp) : ARexp = s match { |
|
480 |
case Nil => r |
|
481 |
case c::s => bders_simp_full(s, bsimp_full(bder(c, r))) |
|
482 |
} |
|
483 |
||
484 |
def blex_simp_full(r: ARexp, s: List[Char]) : Bits = s match { |
|
485 |
case Nil => if (bnullable(r)) bmkeps(r) |
|
486 |
else throw new Exception("Not matched") |
|
487 |
case c::cs => blex_simp_full(bsimp_full(bder(c, r)), cs) |
|
488 |
} |
|
489 |
||
490 |
||
491 |
def blexing_simp_full(r: Rexp, s: String) : Val = |
|
492 |
decode(r, blex_simp_full(internalise(r), s.toList)) |
|
493 |
||
494 |
||
495 |
def btokenise_simp_full(r: Rexp, s: String) = env(blexing_simp_full(r, s)).map(esc2) |
|
496 |
||
497 |
||
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
498 |
|
298 | 499 |
// Testing |
500 |
//============ |
|
501 |
||
502 |
def time[T](code: => T) = { |
|
503 |
val start = System.nanoTime() |
|
504 |
val result = code |
|
505 |
val end = System.nanoTime() |
|
506 |
((end - start)/1.0e9).toString |
|
507 |
//result |
|
508 |
} |
|
509 |
||
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
510 |
def timeR[T](code: => T) = { |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
511 |
val start = System.nanoTime() |
300 | 512 |
for (i <- 1 to 10) code |
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
513 |
val result = code |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
514 |
val end = System.nanoTime() |
300 | 515 |
(result, (end - start)) |
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
516 |
} |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
517 |
|
298 | 518 |
//size: of a Aregx for testing purposes |
519 |
def size(r: Rexp) : Int = r match { |
|
520 |
case ZERO => 1 |
|
521 |
case ONE => 1 |
|
305 | 522 |
case PRED(_,_) => 1 |
298 | 523 |
case SEQ(r1, r2) => 1 + size(r1) + size(r2) |
524 |
case ALTS(rs) => 1 + rs.map(size).sum |
|
525 |
case STAR(r) => 1 + size(r) |
|
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
526 |
case RECD(_, r) => size(r) |
298 | 527 |
} |
528 |
||
529 |
def asize(a: ARexp) = size(erase(a)) |
|
530 |
||
531 |
||
532 |
// Lexing Rules for a Small While Language |
|
533 |
||
534 |
//symbols |
|
305 | 535 |
val SYM = PRED("abcdefghijklmnopqrstuvwxyzABCDEFGHIJKLMNOPQRSTUVWXYZ".contains(_), "SYM") |
298 | 536 |
//digits |
305 | 537 |
val DIGIT = PRED("0123456789".contains(_), "NUM") |
298 | 538 |
//identifiers |
539 |
val ID = SYM ~ (SYM | DIGIT).% |
|
540 |
//numbers |
|
541 |
val NUM = STAR(DIGIT) |
|
542 |
//keywords |
|
543 |
val KEYWORD : Rexp = "skip" | "while" | "do" | "if" | "then" | "else" | "read" | "write" | "true" | "false" |
|
544 |
//semicolons |
|
545 |
val SEMI: Rexp = ";" |
|
546 |
//operators |
|
547 |
val OP: Rexp = ":=" | "==" | "-" | "+" | "*" | "!=" | "<" | ">" | "<=" | ">=" | "%" | "/" |
|
548 |
//whitespaces |
|
549 |
val WHITESPACE = PLUS(" " | "\n" | "\t") |
|
550 |
//parentheses |
|
551 |
val RPAREN: Rexp = ")" |
|
552 |
val LPAREN: Rexp = "(" |
|
553 |
val BEGIN: Rexp = "{" |
|
554 |
val END: Rexp = "}" |
|
555 |
//strings...but probably needs not |
|
556 |
val STRING: Rexp = "\"" ~ SYM.% ~ "\"" |
|
557 |
||
558 |
||
559 |
||
560 |
val WHILE_REGS = (("k" $ KEYWORD) | |
|
561 |
("i" $ ID) | |
|
562 |
("o" $ OP) | |
|
563 |
("n" $ NUM) | |
|
564 |
("s" $ SEMI) | |
|
565 |
("str" $ STRING) | |
|
566 |
("p" $ (LPAREN | RPAREN)) | |
|
567 |
("b" $ (BEGIN | END)) | |
|
568 |
("w" $ WHITESPACE)).% |
|
569 |
||
570 |
// Bigger Tests |
|
571 |
//============== |
|
572 |
||
573 |
||
574 |
val fib_prog = """ |
|
575 |
write "Fib"; |
|
576 |
read n; |
|
577 |
minus1 := 0; |
|
578 |
minus2 := 1; |
|
579 |
while n > 0 do { |
|
580 |
temp := minus2; |
|
581 |
minus2 := minus1 + minus2; |
|
582 |
minus1 := temp; |
|
583 |
n := n - 1 |
|
584 |
}; |
|
585 |
write "Result"; |
|
586 |
write minus2 |
|
587 |
""" |
|
588 |
||
589 |
||
305 | 590 |
|
298 | 591 |
|
305 | 592 |
profile_simp = 0 |
593 |
profile_bsimp = 0 |
|
594 |
profile_flats = 0 |
|
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
595 |
|
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
596 |
println("Original " + size(WHILE_REGS)) |
305 | 597 |
println("Size Bit " + asize(bders_simp((fib_prog * 20).toList, internalise(WHILE_REGS))) + " " + profile_bsimp + " + " + profile_flats) |
598 |
println("Size Old " + size(ders_simp((fib_prog * 20).toList, WHILE_REGS)) + " " + profile_simp) |
|
298 | 599 |
|
305 | 600 |
System.exit(0) |
300 | 601 |
|
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
602 |
println("Internal sizes test OK or strange") |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
603 |
|
300 | 604 |
def perc(p1: Double, p2: Double) : String = |
605 |
f"${(((p1 - p2) / p2) * 100.0) }%5.0f" + "%" |
|
606 |
||
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
607 |
def ders_test(n: Int, s: List[Char], r: Rexp, a: ARexp) : (Rexp, ARexp) = s match { |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
608 |
case Nil => (r, a) |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
609 |
case c::s => { |
300 | 610 |
// derivative |
611 |
val (rd1, tr1) = timeR(der(c, r)) |
|
612 |
val (ad1, ta1) = timeR(bder(c, a)) |
|
613 |
val trs1 = f"${tr1}%.5f" |
|
614 |
val tas1 = f"${ta1}%.5f" |
|
615 |
if (tr1 < ta1) println(s"Time strange der (step) ${n} ${perc(ta1, tr1)} sizes der ${size(rd1)} ${asize(ad1)}") |
|
616 |
//simplification |
|
617 |
val (rd, tr) = timeR(simp(rd1)._1) |
|
618 |
val (ad, ta) = timeR(bsimp(ad1)) |
|
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
619 |
val trs = f"${tr}%.5f" |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
620 |
val tas = f"${ta}%.5f" |
300 | 621 |
//full simplification |
622 |
val (adf, taf) = timeR(bsimp_full(ad1)) |
|
623 |
if (tr < ta) println(s"Time strange simp (step) ${n} ${perc(ta, tr)} sizes simp ${size(rd)} ${asize(ad)}") |
|
624 |
if (n == 1749 || n == 1734) { |
|
625 |
println{s"Aregex before bder (size: ${asize(a)})\n ${string(erase(a))}"} |
|
626 |
println{s"Aregex after bder (size: ${asize(ad1)})\n ${string(erase(ad1))}"} |
|
627 |
println{s"Aregex after bsimp (size: ${asize(ad)})\n ${string(erase(ad))}"} |
|
628 |
println{s"Aregex after bsimp_full (size: ${asize(adf)})\n ${string(erase(adf))}"} |
|
629 |
} |
|
299
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
630 |
ders_test(n + 1, s, rd, ad) |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
631 |
} |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
632 |
} |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
633 |
|
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
634 |
val prg = (fib_prog * 10).toList |
cae7eab03018
added some timing and size tests when doing the derivatives
Christian Urban <urbanc@in.tum.de>
parents:
298
diff
changeset
|
635 |
ders_test(0, prg, WHILE_REGS, internalise(WHILE_REGS)) |
298 | 636 |
|
637 |