src/HOL/ex/SML_Quickcheck_Examples.thy
author wenzelm
Sat, 09 Jul 2011 21:53:27 +0200
changeset 43721 fad8634cee62
parent 43686 bc7d63c7fd6f
permissions -rw-r--r--
echo prover input via raw_messages, for improved protocol tracing;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42161
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/ex/SML_Quickcheck_Examples.thy
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     2
    Author:     Stefan Berghofer, Lukas Bulwahn
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     3
    Copyright   2011 TU Muenchen
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     4
*)
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     5
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     6
theory SML_Quickcheck_Examples
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     7
imports "~~/src/HOL/Library/SML_Quickcheck"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     8
begin
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
     9
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    10
text {*
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    11
This is a regression test for the 'quickcheck' counterexample generator based on the
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    12
SML code generator.
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    13
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    14
The counterexample generator has been superseded by counterexample generators based on
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    15
the generic code generator.
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    16
*}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    17
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    18
declare [[quickcheck_finite_types = false]]
43686
bc7d63c7fd6f tuning options to avoid spurious isabelle test failures
bulwahn
parents: 42161
diff changeset
    19
declare [[quickcheck_timeout = 600.0]]
42161
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    20
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    21
subsection {* Lists *}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    22
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    23
theorem "map g (map f xs) = map (g o f) xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    24
  quickcheck[SML, expect = no_counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    25
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    26
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    27
theorem "map g (map f xs) = map (f o g) xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    28
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    29
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    30
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    31
theorem "rev (xs @ ys) = rev ys @ rev xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    32
  quickcheck[SML, expect = no_counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    33
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    34
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    35
theorem "rev (xs @ ys) = rev xs @ rev ys"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    36
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    37
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    38
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    39
theorem "rev (rev xs) = xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    40
  quickcheck[SML, expect = no_counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    41
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    42
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    43
theorem "rev xs = xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    44
  quickcheck[tester = SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    45
oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    46
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    47
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    48
text {* An example involving functions inside other data structures *}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    49
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    50
primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    51
  "app [] x = x"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    52
  | "app (f # fs) x = app fs (f x)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    53
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    54
lemma "app (fs @ gs) x = app gs (app fs x)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    55
  quickcheck[SML, expect = no_counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    56
  by (induct fs arbitrary: x) simp_all
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    57
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    58
lemma "app (fs @ gs) x = app fs (app gs x)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    59
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    60
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    61
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    62
primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    63
  "occurs a [] = 0"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    64
  | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    65
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    66
primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    67
  "del1 a [] = []"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    68
  | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    69
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    70
text {* A lemma, you'd think to be true from our experience with delAll *}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    71
lemma "Suc (occurs a (del1 a xs)) = occurs a xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    72
  -- {* Wrong. Precondition needed.*}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    73
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    74
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    75
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    76
lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    77
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    78
    -- {* Also wrong.*}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    79
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    80
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    81
lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    82
  quickcheck[SML, expect = no_counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    83
  by (induct xs) auto
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    84
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    85
primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    86
  "replace a b [] = []"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    87
  | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) 
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    88
                            else (x#(replace a b xs)))"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    89
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    90
lemma "occurs a xs = occurs b (replace a b xs)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    91
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    92
  -- {* Wrong. Precondition needed.*}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    93
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    94
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    95
lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    96
  quickcheck[SML, expect = no_counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    97
  by (induct xs) simp_all
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    98
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
    99
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   100
subsection {* Trees *}
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   101
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   102
datatype 'a tree = Twig |  Leaf 'a | Branch "'a tree" "'a tree"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   103
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   104
primrec leaves :: "'a tree \<Rightarrow> 'a list" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   105
  "leaves Twig = []"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   106
  | "leaves (Leaf a) = [a]"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   107
  | "leaves (Branch l r) = (leaves l) @ (leaves r)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   108
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   109
primrec plant :: "'a list \<Rightarrow> 'a tree" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   110
  "plant [] = Twig "
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   111
  | "plant (x#xs) = Branch (Leaf x) (plant xs)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   112
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   113
primrec mirror :: "'a tree \<Rightarrow> 'a tree" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   114
  "mirror (Twig) = Twig "
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   115
  | "mirror (Leaf a) = Leaf a "
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   116
  | "mirror (Branch l r) = Branch (mirror r) (mirror l)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   117
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   118
theorem "plant (rev (leaves xt)) = mirror xt"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   119
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   120
    --{* Wrong! *} 
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   121
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   122
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   123
theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   124
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   125
    --{* Wrong! *} 
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   126
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   127
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   128
datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   129
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   130
primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   131
  "inOrder (Tip a)= [a]"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   132
  | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   133
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   134
primrec root :: "'a ntree \<Rightarrow> 'a" where
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   135
  "root (Tip a) = a"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   136
  | "root (Node f x y) = f"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   137
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   138
theorem "hd (inOrder xt) = root xt"
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   139
  quickcheck[SML, expect = counterexample]
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   140
  --{* Wrong! *} 
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   141
  oops
d1b39536e1fb adding a regression test for SML_Quickcheck; putting SML_Quickcheck back in shape
bulwahn
parents:
diff changeset
   142
43686
bc7d63c7fd6f tuning options to avoid spurious isabelle test failures
bulwahn
parents: 42161
diff changeset
   143
end