src/HOL/Predicate_Compile_Examples/Reg_Exp_Example.thy
author blanchet
Tue, 09 Sep 2014 20:51:36 +0200
changeset 58249 180f1b3508ed
parent 56927 4044a7d1720f
child 58310 91ea607a34d8
permissions -rw-r--r--
use 'datatype_new' (soon to be renamed 'datatype') in Isabelle's libraries
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
41956
c15ef1b85035 modernized imports (untested!?);
wenzelm
parents: 41413
diff changeset
     2
imports
c15ef1b85035 modernized imports (untested!?);
wenzelm
parents: 41413
diff changeset
     3
  "~~/src/HOL/Library/Predicate_Compile_Quickcheck"
c15ef1b85035 modernized imports (untested!?);
wenzelm
parents: 41413
diff changeset
     4
  "~~/src/HOL/Library/Code_Prolog"
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
     5
begin
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
     6
54703
499f92dc6e45 more antiquotations;
wenzelm
parents: 52666
diff changeset
     7
text {* An example from the experiments from SmallCheck (@{url "http://www.cs.york.ac.uk/fp/smallcheck/"}) *}
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
     8
text {* The example (original in Haskell) was imported with Haskabelle and
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
     9
  manually slightly adapted.
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    10
*}
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    11
 
58249
180f1b3508ed use 'datatype_new' (soon to be renamed 'datatype') in Isabelle's libraries
blanchet
parents: 56927
diff changeset
    12
datatype_new Nat = Zer
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    13
             | Suc Nat
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    14
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    15
fun sub :: "Nat \<Rightarrow> Nat \<Rightarrow> Nat"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    16
where
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    17
  "sub x y = (case y of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    18
                 Zer \<Rightarrow> x
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    19
               | Suc y' \<Rightarrow> case x of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    20
                                         Zer \<Rightarrow> Zer
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    21
                                       | Suc x' \<Rightarrow> sub x' y')"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    22
58249
180f1b3508ed use 'datatype_new' (soon to be renamed 'datatype') in Isabelle's libraries
blanchet
parents: 56927
diff changeset
    23
datatype_new Sym = N0
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    24
             | N1 Sym
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    25
58249
180f1b3508ed use 'datatype_new' (soon to be renamed 'datatype') in Isabelle's libraries
blanchet
parents: 56927
diff changeset
    26
datatype_new RE = Sym Sym
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    27
            | Or RE RE
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    28
            | Seq RE RE
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    29
            | And RE RE
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    30
            | Star RE
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    31
            | Empty
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    32
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    33
 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    34
function accepts :: "RE \<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" and 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    36
    seqSplit'' :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool" and 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    37
    seqSplit' :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    38
where
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    39
  "accepts re ss = (case re of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    40
                       Sym n \<Rightarrow> case ss of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    41
                                              Nil \<Rightarrow> False
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    42
                                            | (n' # ss') \<Rightarrow> n = n' & List.null ss'
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    43
                     | Or re1 re2 \<Rightarrow> accepts re1 ss | accepts re2 ss
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    44
                     | Seq re1 re2 \<Rightarrow> seqSplit re1 re2 Nil ss
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    45
                     | And re1 re2 \<Rightarrow> accepts re1 ss & accepts re2 ss
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    46
                     | Star re' \<Rightarrow> case ss of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    47
                                                 Nil \<Rightarrow> True
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    48
                                               | (s # ss') \<Rightarrow> seqSplit re' re [s] ss'
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    49
                     | Empty \<Rightarrow> List.null ss)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    50
| "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
    51
| "seqSplit'' re1 re2 ss2 ss = (accepts re1 ss2 & accepts re2 ss)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    52
| "seqSplit' re1 re2 ss2 ss = (case ss of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    53
                                  Nil \<Rightarrow> False
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    54
                                | (n # ss') \<Rightarrow> seqSplit re1 re2 (ss2 @ [n]) ss')"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    55
by pat_completeness auto
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    56
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    57
termination
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    58
  sorry
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    59
 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    60
fun rep :: "Nat \<Rightarrow> RE \<Rightarrow> RE"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    61
where
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    62
  "rep n re = (case n of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    63
                  Zer \<Rightarrow> Empty
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    64
                | Suc n' \<Rightarrow> Seq re (rep n' re))"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    65
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    66
 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    67
fun repMax :: "Nat \<Rightarrow> RE \<Rightarrow> RE"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    68
where
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    69
  "repMax n re = (case n of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    70
                     Zer \<Rightarrow> Empty
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    71
                   | Suc n' \<Rightarrow> Or (rep n re) (repMax n' re))"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    72
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    73
 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    74
fun repInt' :: "Nat \<Rightarrow> Nat \<Rightarrow> RE \<Rightarrow> RE"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    75
where
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    76
  "repInt' n k re = (case k of
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    77
                        Zer \<Rightarrow> rep n re
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    78
                      | Suc k' \<Rightarrow> Or (rep n re) (repInt' (Suc n) k' re))"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    79
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    80
 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    81
fun repInt :: "Nat \<Rightarrow> Nat \<Rightarrow> RE \<Rightarrow> RE"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    82
where
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    83
  "repInt n k re = repInt' n (sub k n) re"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    84
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    85
 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    86
fun prop_regex :: "Nat * Nat * RE * RE * Sym list \<Rightarrow> bool"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    87
where
52666
391913d17d15 tuned line length;
wenzelm
parents: 45451
diff changeset
    88
  "prop_regex (n, (k, (p, (q, s)))) =
391913d17d15 tuned line length;
wenzelm
parents: 45451
diff changeset
    89
    ((accepts (repInt n k (And p q)) s) = (accepts (And (repInt n k p) (repInt n k q)) s))"
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    90
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    91
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    92
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    93
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
    94
(*nitpick
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
    95
quickcheck[tester = random, iterations = 10000]
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
    96
quickcheck[tester = predicate_compile_wo_ff]
45451
74515e8e6046 renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
bulwahn
parents: 43974
diff changeset
    97
quickcheck[tester = predicate_compile_ff_fs]*)
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    98
oops
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
    99
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   100
52666
391913d17d15 tuned line length;
wenzelm
parents: 45451
diff changeset
   101
setup {*
391913d17d15 tuned line length;
wenzelm
parents: 45451
diff changeset
   102
  Context.theory_map
391913d17d15 tuned line length;
wenzelm
parents: 45451
diff changeset
   103
    (Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals)))
391913d17d15 tuned line length;
wenzelm
parents: 45451
diff changeset
   104
*}
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   105
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   106
setup {* Code_Prolog.map_code_options (K 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   107
  {ensure_groundness = true,
39800
17e29ddd538e adapting manual configuration in examples
bulwahn
parents: 39725
diff changeset
   108
   limit_globally = NONE,
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   109
   limited_types = [(@{typ Sym}, 0), (@{typ "Sym list"}, 2), (@{typ RE}, 6)],
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   110
   limited_predicates = [(["repIntPa"], 2), (["repP"], 2), (["subP"], 0),
39725
4f4dd39a1e58 adopting example
bulwahn
parents: 39463
diff changeset
   111
    (["accepts", "acceptsaux", "seqSplit", "seqSplita", "seqSplitaux", "seqSplitb"], 25)],
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   112
   replacing =
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   113
     [(("repP", "limited_repP"), "lim_repIntPa"),
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   114
      (("subP", "limited_subP"), "repIntP"),
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   115
      (("repIntPa", "limited_repIntPa"), "repIntP"),
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   116
      (("accepts", "limited_accepts"), "quickcheck")],  
39463
7ce0ed8dc4d6 adapting examples
bulwahn
parents: 39196
diff changeset
   117
   manual_reorder = []}) *}
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   118
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   119
text {* Finding the counterexample still seems out of reach for the
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   120
 prolog-style generation. *}
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   121
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   122
lemma "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
45451
74515e8e6046 renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
bulwahn
parents: 43974
diff changeset
   123
quickcheck[exhaustive]
74515e8e6046 renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
bulwahn
parents: 43974
diff changeset
   124
quickcheck[tester = random, iterations = 1000]
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
   125
(*quickcheck[tester = predicate_compile_wo_ff]*)
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
   126
(*quickcheck[tester = predicate_compile_ff_fs, iterations = 1]*)
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
   127
(*quickcheck[tester = prolog, iterations = 1, size = 1]*)
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   128
oops
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   129
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   130
section {* Given a partial solution *}
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   131
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   132
lemma
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   133
(*  assumes "n = Zer"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   134
  assumes "k = Suc (Suc Zer)"*)
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   135
  assumes "p = Sym N0"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   136
  assumes "q = Seq (Sym N0) (Sym N0)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   137
(*  assumes "s = [N0, N0]"*)
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   138
  shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
   139
(*quickcheck[tester = predicate_compile_wo_ff, iterations = 1]*)
43974
30f4d346b204 removing expectations from quickcheck example
bulwahn
parents: 43953
diff changeset
   140
quickcheck[tester = prolog, iterations = 1, size = 1]
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   141
oops
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   142
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   143
section {* Checking the counterexample *}
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   144
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   145
definition a_sol
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   146
where
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   147
  "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
   148
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   149
lemma 
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   150
  assumes "n = Zer"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   151
  assumes "k = Suc (Suc Zer)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   152
  assumes "p = Sym N0"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   153
  assumes "q = Seq (Sym N0) (Sym N0)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   154
  assumes "s = [N0, N0]"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   155
  shows "accepts (repInt n k (And p q)) s --> accepts (And (repInt n k p) (repInt n k q)) s"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
   156
(*quickcheck[tester = predicate_compile_wo_ff]*)
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   157
oops
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   158
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   159
lemma
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   160
  assumes "n = Zer"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   161
  assumes "k = Suc (Suc Zer)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   162
  assumes "p = Sym N0"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   163
  assumes "q = Seq (Sym N0) (Sym N0)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   164
  assumes "s = [N0, N0]"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   165
  shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39800
diff changeset
   166
(*quickcheck[tester = predicate_compile_wo_ff, iterations = 1, expect = counterexample]*)
43974
30f4d346b204 removing expectations from quickcheck example
bulwahn
parents: 43953
diff changeset
   167
quickcheck[tester = prolog, iterations = 1, size = 1]
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   168
oops
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   169
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   170
lemma
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   171
  assumes "n = Zer"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   172
  assumes "k = Suc (Suc Zer)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   173
  assumes "p = Sym N0"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   174
  assumes "q = Seq (Sym N0) (Sym N0)"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   175
  assumes "s = [N0, N0]"
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   176
shows "prop_regex (n, (k, (p, (q, s))))"
45451
74515e8e6046 renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
bulwahn
parents: 43974
diff changeset
   177
quickcheck[tester = smart_exhaustive, depth = 30]
43974
30f4d346b204 removing expectations from quickcheck example
bulwahn
parents: 43953
diff changeset
   178
quickcheck[tester = prolog]
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   179
oops
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   180
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   181
lemma "prop_regex a_sol"
43974
30f4d346b204 removing expectations from quickcheck example
bulwahn
parents: 43953
diff changeset
   182
quickcheck[tester = random]
45451
74515e8e6046 renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
bulwahn
parents: 43974
diff changeset
   183
quickcheck[tester = smart_exhaustive]
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   184
oops
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   185
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 54703
diff changeset
   186
value "prop_regex a_sol"
39188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   187
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   188
cd6558ed65d7 adding the Reg_Exp example
bulwahn
parents:
diff changeset
   189
end