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