| 39189 |      1 | theory Reg_Exp_Example
 | 
| 41956 |      2 | imports
 | 
|  |      3 |   "~~/src/HOL/Library/Predicate_Compile_Quickcheck"
 | 
|  |      4 |   "~~/src/HOL/Library/Code_Prolog"
 | 
| 39188 |      5 | begin
 | 
|  |      6 | 
 | 
|  |      7 | text {* An example from the experiments from SmallCheck (http://www.cs.york.ac.uk/fp/smallcheck/) *}
 | 
|  |      8 | text {* The example (original in Haskell) was imported with Haskabelle and
 | 
|  |      9 |   manually slightly adapted.
 | 
|  |     10 | *}
 | 
|  |     11 |  
 | 
|  |     12 | datatype Nat = Zer
 | 
|  |     13 |              | Suc Nat
 | 
|  |     14 | 
 | 
|  |     15 | fun sub :: "Nat \<Rightarrow> Nat \<Rightarrow> Nat"
 | 
|  |     16 | where
 | 
|  |     17 |   "sub x y = (case y of
 | 
|  |     18 |                  Zer \<Rightarrow> x
 | 
|  |     19 |                | Suc y' \<Rightarrow> case x of
 | 
|  |     20 |                                          Zer \<Rightarrow> Zer
 | 
|  |     21 |                                        | Suc x' \<Rightarrow> sub x' y')"
 | 
|  |     22 | 
 | 
|  |     23 | datatype Sym = N0
 | 
|  |     24 |              | N1 Sym
 | 
|  |     25 | 
 | 
|  |     26 | datatype RE = Sym Sym
 | 
|  |     27 |             | Or RE RE
 | 
|  |     28 |             | Seq RE RE
 | 
|  |     29 |             | And RE RE
 | 
|  |     30 |             | Star RE
 | 
|  |     31 |             | Empty
 | 
|  |     32 | 
 | 
|  |     33 |  
 | 
|  |     34 | function accepts :: "RE \<Rightarrow> Sym list \<Rightarrow> bool" and 
 | 
|  |     35 |     seqSplit :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool" and 
 | 
|  |     36 |     seqSplit'' :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool" and 
 | 
|  |     37 |     seqSplit' :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool"
 | 
|  |     38 | where
 | 
|  |     39 |   "accepts re ss = (case re of
 | 
|  |     40 |                        Sym n \<Rightarrow> case ss of
 | 
|  |     41 |                                               Nil \<Rightarrow> False
 | 
|  |     42 |                                             | (n' # ss') \<Rightarrow> n = n' & List.null ss'
 | 
|  |     43 |                      | Or re1 re2 \<Rightarrow> accepts re1 ss | accepts re2 ss
 | 
|  |     44 |                      | Seq re1 re2 \<Rightarrow> seqSplit re1 re2 Nil ss
 | 
|  |     45 |                      | And re1 re2 \<Rightarrow> accepts re1 ss & accepts re2 ss
 | 
|  |     46 |                      | Star re' \<Rightarrow> case ss of
 | 
|  |     47 |                                                  Nil \<Rightarrow> True
 | 
|  |     48 |                                                | (s # ss') \<Rightarrow> seqSplit re' re [s] ss'
 | 
|  |     49 |                      | Empty \<Rightarrow> List.null ss)"
 | 
|  |     50 | | "seqSplit re1 re2 ss2 ss = (seqSplit'' re1 re2 ss2 ss | seqSplit' re1 re2 ss2 ss)"
 | 
|  |     51 | | "seqSplit'' re1 re2 ss2 ss = (accepts re1 ss2 & accepts re2 ss)"
 | 
|  |     52 | | "seqSplit' re1 re2 ss2 ss = (case ss of
 | 
|  |     53 |                                   Nil \<Rightarrow> False
 | 
|  |     54 |                                 | (n # ss') \<Rightarrow> seqSplit re1 re2 (ss2 @ [n]) ss')"
 | 
|  |     55 | by pat_completeness auto
 | 
|  |     56 | 
 | 
|  |     57 | termination
 | 
|  |     58 |   sorry
 | 
|  |     59 |  
 | 
|  |     60 | fun rep :: "Nat \<Rightarrow> RE \<Rightarrow> RE"
 | 
|  |     61 | where
 | 
|  |     62 |   "rep n re = (case n of
 | 
|  |     63 |                   Zer \<Rightarrow> Empty
 | 
|  |     64 |                 | Suc n' \<Rightarrow> Seq re (rep n' re))"
 | 
|  |     65 | 
 | 
|  |     66 |  
 | 
|  |     67 | fun repMax :: "Nat \<Rightarrow> RE \<Rightarrow> RE"
 | 
|  |     68 | where
 | 
|  |     69 |   "repMax n re = (case n of
 | 
|  |     70 |                      Zer \<Rightarrow> Empty
 | 
|  |     71 |                    | Suc n' \<Rightarrow> Or (rep n re) (repMax n' re))"
 | 
|  |     72 | 
 | 
|  |     73 |  
 | 
|  |     74 | fun repInt' :: "Nat \<Rightarrow> Nat \<Rightarrow> RE \<Rightarrow> RE"
 | 
|  |     75 | where
 | 
|  |     76 |   "repInt' n k re = (case k of
 | 
|  |     77 |                         Zer \<Rightarrow> rep n re
 | 
|  |     78 |                       | Suc k' \<Rightarrow> Or (rep n re) (repInt' (Suc n) k' re))"
 | 
|  |     79 | 
 | 
|  |     80 |  
 | 
|  |     81 | fun repInt :: "Nat \<Rightarrow> Nat \<Rightarrow> RE \<Rightarrow> RE"
 | 
|  |     82 | where
 | 
|  |     83 |   "repInt n k re = repInt' n (sub k n) re"
 | 
|  |     84 | 
 | 
|  |     85 |  
 | 
|  |     86 | fun prop_regex :: "Nat * Nat * RE * RE * Sym list \<Rightarrow> bool"
 | 
|  |     87 | where
 | 
|  |     88 |   "prop_regex (n, (k, (p, (q, s)))) = ((accepts (repInt n k (And p q)) s) = (accepts (And (repInt n k p) (repInt n k q)) s))"
 | 
|  |     89 | 
 | 
|  |     90 | 
 | 
|  |     91 | 
 | 
|  |     92 | lemma "accepts (repInt n k (And p q)) s --> accepts (And (repInt n k p) (repInt n k q)) s"
 | 
|  |     93 | (*nitpick
 | 
| 40924 |     94 | quickcheck[tester = random, iterations = 10000]
 | 
|  |     95 | quickcheck[tester = predicate_compile_wo_ff]
 | 
|  |     96 | quickcheck[tester = predicate_compile_ff_fs]
 | 
| 39188 |     97 | oops*)
 | 
|  |     98 | oops
 | 
|  |     99 | 
 | 
|  |    100 | 
 | 
| 39463 |    101 | setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
 | 
| 39188 |    102 | 
 | 
|  |    103 | setup {* Code_Prolog.map_code_options (K 
 | 
|  |    104 |   {ensure_groundness = true,
 | 
| 39800 |    105 |    limit_globally = NONE,
 | 
| 39188 |    106 |    limited_types = [(@{typ Sym}, 0), (@{typ "Sym list"}, 2), (@{typ RE}, 6)],
 | 
|  |    107 |    limited_predicates = [(["repIntPa"], 2), (["repP"], 2), (["subP"], 0),
 | 
| 39725 |    108 |     (["accepts", "acceptsaux", "seqSplit", "seqSplita", "seqSplitaux", "seqSplitb"], 25)],
 | 
| 39188 |    109 |    replacing =
 | 
|  |    110 |      [(("repP", "limited_repP"), "lim_repIntPa"),
 | 
|  |    111 |       (("subP", "limited_subP"), "repIntP"),
 | 
|  |    112 |       (("repIntPa", "limited_repIntPa"), "repIntP"),
 | 
|  |    113 |       (("accepts", "limited_accepts"), "quickcheck")],  
 | 
| 39463 |    114 |    manual_reorder = []}) *}
 | 
| 39188 |    115 | 
 | 
|  |    116 | text {* Finding the counterexample still seems out of reach for the
 | 
|  |    117 |  prolog-style generation. *}
 | 
|  |    118 | 
 | 
|  |    119 | lemma "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
 | 
| 40924 |    120 | quickcheck[tester = random, iterations = 100000, expect = counterexample]
 | 
|  |    121 | (*quickcheck[tester = predicate_compile_wo_ff]*)
 | 
|  |    122 | (*quickcheck[tester = predicate_compile_ff_fs, iterations = 1]*)
 | 
|  |    123 | (*quickcheck[tester = prolog, iterations = 1, size = 1]*)
 | 
| 39188 |    124 | oops
 | 
|  |    125 | 
 | 
|  |    126 | section {* Given a partial solution *}
 | 
|  |    127 | 
 | 
|  |    128 | lemma
 | 
|  |    129 | (*  assumes "n = Zer"
 | 
|  |    130 |   assumes "k = Suc (Suc Zer)"*)
 | 
|  |    131 |   assumes "p = Sym N0"
 | 
|  |    132 |   assumes "q = Seq (Sym N0) (Sym N0)"
 | 
|  |    133 | (*  assumes "s = [N0, N0]"*)
 | 
|  |    134 |   shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
 | 
| 40924 |    135 | (*quickcheck[tester = predicate_compile_wo_ff, iterations = 1]*)
 | 
|  |    136 | quickcheck[tester = prolog, iterations = 1, size = 1, expect = counterexample]
 | 
| 39188 |    137 | oops
 | 
|  |    138 | 
 | 
|  |    139 | section {* Checking the counterexample *}
 | 
|  |    140 | 
 | 
|  |    141 | definition a_sol
 | 
|  |    142 | where
 | 
|  |    143 |   "a_sol = (Zer, (Suc (Suc Zer), (Sym N0, (Seq (Sym N0) (Sym N0), [N0, N0]))))"
 | 
|  |    144 | 
 | 
|  |    145 | lemma 
 | 
|  |    146 |   assumes "n = Zer"
 | 
|  |    147 |   assumes "k = Suc (Suc Zer)"
 | 
|  |    148 |   assumes "p = Sym N0"
 | 
|  |    149 |   assumes "q = Seq (Sym N0) (Sym N0)"
 | 
|  |    150 |   assumes "s = [N0, N0]"
 | 
|  |    151 |   shows "accepts (repInt n k (And p q)) s --> accepts (And (repInt n k p) (repInt n k q)) s"
 | 
| 40924 |    152 | (*quickcheck[tester = predicate_compile_wo_ff]*)
 | 
| 39188 |    153 | oops
 | 
|  |    154 | 
 | 
|  |    155 | lemma
 | 
|  |    156 |   assumes "n = Zer"
 | 
|  |    157 |   assumes "k = Suc (Suc Zer)"
 | 
|  |    158 |   assumes "p = Sym N0"
 | 
|  |    159 |   assumes "q = Seq (Sym N0) (Sym N0)"
 | 
|  |    160 |   assumes "s = [N0, N0]"
 | 
|  |    161 |   shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
 | 
| 40924 |    162 | (*quickcheck[tester = predicate_compile_wo_ff, iterations = 1, expect = counterexample]*)
 | 
|  |    163 | quickcheck[tester = prolog, iterations = 1, size = 1, expect = counterexample]
 | 
| 39188 |    164 | oops
 | 
|  |    165 | 
 | 
|  |    166 | lemma
 | 
|  |    167 |   assumes "n = Zer"
 | 
|  |    168 |   assumes "k = Suc (Suc Zer)"
 | 
|  |    169 |   assumes "p = Sym N0"
 | 
|  |    170 |   assumes "q = Seq (Sym N0) (Sym N0)"
 | 
|  |    171 |   assumes "s = [N0, N0]"
 | 
|  |    172 | shows "prop_regex (n, (k, (p, (q, s))))"
 | 
| 40924 |    173 | quickcheck[tester = predicate_compile_wo_ff]
 | 
|  |    174 | quickcheck[tester = prolog, expect = counterexample]
 | 
| 39188 |    175 | oops
 | 
|  |    176 | 
 | 
|  |    177 | lemma "prop_regex a_sol"
 | 
| 40924 |    178 | quickcheck[tester = random, expect = counterexample]
 | 
|  |    179 | quickcheck[tester = predicate_compile_ff_fs]
 | 
| 39188 |    180 | oops
 | 
|  |    181 | 
 | 
|  |    182 | value [code] "prop_regex a_sol"
 | 
|  |    183 | 
 | 
|  |    184 | 
 | 
|  |    185 | end
 |