# HG changeset patch # User Chengsong # Date 1655999268 -3600 # Node ID 333013923c5a500f8d87f569b1c2ec839be85d1d # Parent b672be21ffac4ddfb98d8c4a7729cc244e25e64f more changes to blexersimp.thy diff -r b672be21ffac -r 333013923c5a RegexExplosionPlot/evilForBsimp.sc --- a/RegexExplosionPlot/evilForBsimp.sc Thu Jun 23 16:10:04 2022 +0100 +++ b/RegexExplosionPlot/evilForBsimp.sc Thu Jun 23 16:47:48 2022 +0100 @@ -1027,7 +1027,7 @@ for(n <- 1 to 12){ val STARREG = n_astar_alts(n) val iMax = (lcm((1 to n).toList)) - for(i <- iMax to iMax ){// 100, 400, 800, 840, 841, 900 + for(i <- 1 to iMax + 1 ){// 100, 400, 800, 840, 841, 900 val prog0 = "a" * i //println(s"test: $prog0") print(n) diff -r b672be21ffac -r 333013923c5a thys2/blexer2.sc --- a/thys2/blexer2.sc Thu Jun 23 16:10:04 2022 +0100 +++ b/thys2/blexer2.sc Thu Jun 23 16:47:48 2022 +0100 @@ -944,14 +944,6 @@ rs1.forall(rs2.contains) def prune6(acc: Set[Rexp], r: ARexp, ctx: List[Rexp] = Nil) : ARexp = { - // if (erase(r) == ONE && acc != Set()) - // { - // println("ctx") - // rsprint(ctx) - // println("acc") - // rsprint(acc) - // println("acc-end") - // } if (ABIncludedByC(a = r, b = ctx, c = acc, f = attachCtx, subseteqPred = rs1_subseteq_rs2)) { diff -r b672be21ffac -r 333013923c5a thys3/BlexerSimp.thy --- a/thys3/BlexerSimp.thy Thu Jun 23 16:10:04 2022 +0100 +++ b/thys3/BlexerSimp.thy Thu Jun 23 16:47:48 2022 +0100 @@ -614,22 +614,6 @@ -fun bsimpStrong6 :: "arexp \ arexp" - where - "bsimpStrong6 (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimpStrong6 r1) (bsimpStrong6 r2)" -| "bsimpStrong6 (AALTs bs1 rs) = bsimp_AALTs bs1 (distinctWith (flts (map bsimpStrong6 rs)) eq1 {}) " -| "bsimpStrong6 r = r" - - -fun - bdersStrong6 :: "arexp \ string \ arexp" -where - "bdersStrong6 r [] = r" -| "bdersStrong6 r (c # s) = bdersStrong6 (bsimpStrong6 (bder c r)) s" - -definition blexerStrong where - "blexerStrong r s \ if bnullable (bdersStrong6 (intern r) s) then - decode (bmkeps (bdersStrong6 (intern r) s)) r else None" fun ABIncludedByC :: "'a \ 'b \ 'c \ ('a \ 'b \ 'c) \ ('c \ 'c \ bool) \ bool" where "ABIncludedByC a b c f subseteqPred = subseteqPred (f a b) c" @@ -651,52 +635,119 @@ fun attachCtx :: "arexp \ rexp list \ rexp set" where "attachCtx r ctx = set (turnIntoTerms (regConcat (erase r) ctx))" -(* -def regConcat(acc: Rexp, rs: List[Rexp]) : Rexp = rs match { - case Nil => acc - case r :: rs1 => - // if(acc == ONE) - // regConcat(r, rs1) - // else - regConcat(SEQ(acc, r), rs1) -} -def attachCtx(r: ARexp, ctx: List[Rexp]) : Set[Rexp] = { - turnIntoTerms((regConcat(erase(r), ctx))) - .toSet -} - -def attachCtxcc(r: Rexp, ctx: List[Rexp]) : Set[Rexp] = - turnIntoTerms(regConcat(r, ctx)).toSet - -def ABIncludedByC[A, B, C](a: A, b: B, c: C, f: (A, B) => C, -subseteqPred: (C, C) => Boolean) : Boolean = { - subseteqPred(f(a, b), c) -} -def rs1_subseteq_rs2(rs1: Set[Rexp], rs2: Set[Rexp]) : Boolean = - //rs1 \subseteq rs2 - rs1.forall(rs2.contains) - - -} -*) fun rs1_subseteq_rs2 :: "rexp set \ rexp set \ bool" where "rs1_subseteq_rs2 rs1 rs2 = (rs1 \ rs2)" +fun notZero :: "arexp \ bool" where + "notZero AZERO = True" +| "notZero _ = False" + fun prune6 :: "rexp set \ arexp \ rexp list \ arexp" where "prune6 acc a ctx = (if (ABIncludedByC a ctx acc attachCtx rs1_subseteq_rs2) then AZERO else (case a of (ASEQ bs r1 r2) \ bsimp_ASEQ bs (prune6 acc r1 (erase r2 # ctx)) r2 - | AALTs bs rs0 \ bsimp_AALTs bs (map (\r. filter notZero (prune6 acc r ctx)) rs0)) )" + | AALTs bs rs0 \ bsimp_AALTs bs (filter notZero (map (\r.(prune6 acc r ctx)) rs0))) )" + + +(* +def prune6(acc: Set[Rexp], r: ARexp, ctx: List[Rexp] = Nil) : ARexp = { + if (ABIncludedByC(a = r, b = ctx, c = acc, + f = attachCtx, subseteqPred = rs1_subseteq_rs2)) + { + AZERO + } + else{ + r match { + case ASEQ(bs, r1, r2) => (prune6(acc, r1, erase(r2) :: ctx)) match{ + case AZERO => + AZERO + case AONE(bs1) => //when r1 becomes 1, chances to further prune r2 + prune6(acc, fuse(bs1, r2), ctx) + case r1p => + ASEQ(bs, r1p, r2) + } + case AALTS(bs, rs0) => + rs0.map(r => prune6(acc, r, ctx)).filter(_ != AZERO) match { + case Nil => + AZERO + case r :: Nil => + fuse(bs, r) + case rs0p => + AALTS(bs, rs0p) + } + case r => r + } + } +} + + +*) fun dB6 :: "arexp list \ rexp set \ arexp list" where "dB6 [] acc = []" | "dB6 (a # as) acc = (if (erase a \ acc) then (dB6 as acc) - else (let pruned = prune6 acc x in + else (let pruned = prune6 acc a [] in (case pruned of AZERO \ dB6 as acc - |xPrime \ xPrime # (dB6 xs ((turnIntoTerms (erase pruned)) \ acc) ) ) )) " + |xPrime \ xPrime # (dB6 as ( (set (turnIntoTerms (erase pruned))) \ acc) ) ) )) " + + +fun bsimpStrong6 :: "arexp \ arexp" + where + "bsimpStrong6 (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimpStrong6 r1) (bsimpStrong6 r2)" +| "bsimpStrong6 (AALTs bs1 rs) = bsimp_AALTs bs1 (dB6 (flts (map bsimpStrong6 rs)) {}) " +| "bsimpStrong6 r = r" + + +fun + bdersStrong6 :: "arexp \ string \ arexp" +where + "bdersStrong6 r [] = r" +| "bdersStrong6 r (c # s) = bdersStrong6 (bsimpStrong6 (bder c r)) s" + +definition blexerStrong where + "blexerStrong r s \ if bnullable (bdersStrong6 (intern r) s) then + decode (bmkeps (bdersStrong6 (intern r) s)) r else None" + + +lemma centralStrong: + shows "bders_simp r s \* bdersStrong6 r s" +proof(induct s arbitrary: r rule: rev_induct) + case Nil + then show "bders_simp r [] \* bdersStrong6 r []" by simp +next + case (snoc x xs) + have IH: "\r. bders_simp r xs \* bdersStrong6 r xs" by fact + have "bders_simp r (xs @ [x]) = bders_simp (bders_simp r xs) [x]" by (simp add: bders_simp_append) + also have "... \* bders_simp (bdersStrong6 r xs) [x]" using IH + by (simp add: rewrites_preserves_bder) + also have "... \* bders_simp (bders_simp r xs) [x]" using IH + by (simp add: rewrites_to_bsimp) + finally show "bders r (xs @ [x]) \* bders_simp r (xs @ [x])" + by (simp add: bders_simp_append) +qed + +lemma mainStrong: + assumes "bnullable (bders r s)" + shows "bmkeps (bders_simp r s) = bmkeps (bdersStrong6 r s)" +proof - + have "bders_simp r s \* bdersStrong6 r s" sorry + then + show "bmkeps (bders_simp r s) = bmkeps (bdersStrong6 r s)" + + sorry +qed + + + + +theorem blexerStrong_correct : + shows "blexerStrong r s = blexer_simp r s" + + sorry + end