src/HOL/ex/Predicate_Compile_ex.thy
author bulwahn
Sat, 24 Oct 2009 16:55:42 +0200
changeset 33114 4785ef554dcc
parent 33113 0f6e30b87cf1
child 33116 b379ee2cddb1
permissions -rw-r--r--
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31129
d2cead76fca2 split Predicate_Compile examples into separate theory
haftmann
parents: 31123
diff changeset
     1
theory Predicate_Compile_ex
32318
bca7fd849829 improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents: 32317
diff changeset
     2
imports Main Predicate_Compile
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     3
begin
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     4
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     5
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     6
    "even 0"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     7
  | "even n \<Longrightarrow> odd (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     8
  | "odd n \<Longrightarrow> even (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     9
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    10
code_pred (mode: [], [1]) even .
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31514
diff changeset
    11
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
    12
thm odd.equation
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    13
thm even.equation
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    14
(*
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    15
lemma unit: "unit_case f = (\<lambda>x. f)"
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    16
apply (rule ext)
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    17
apply (case_tac x)
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    18
apply (simp only: unit.cases)
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    19
done
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    20
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    21
lemma "even_1 (Suc (Suc n)) = even_1 n"
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    22
thm even.equation(2)
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    23
unfolding even.equation(1)[of "Suc (Suc n)"]
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    24
unfolding odd.equation
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    25
unfolding single_bind
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    26
apply simp
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    27
apply (simp add: unit)
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    28
unfolding bind_single
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    29
apply (rule refl)
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    30
done
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    31
*)
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    32
values "{x. even 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    33
values "{x. odd 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    34
values 10 "{n. even n}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    35
values 10 "{n. odd n}"
31195
12741f23527d added example on ML level
haftmann
parents: 31129
diff changeset
    36
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    37
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
    38
    "append [] xs xs"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
    39
  | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    40
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    41
code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    42
thm append.equation
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    43
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
    44
code_pred (inductify_all) (rpred) append .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    45
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    46
thm append.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
    47
thm append.rpred_equation
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    48
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    49
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    50
values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    51
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0,5]}"
31195
12741f23527d added example on ML level
haftmann
parents: 31129
diff changeset
    52
33111
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
    53
value [code] "Predicate.the (append_1_2 [0::int, 1, 2] [3, 4, 5])"
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
    54
value [code] "Predicate.the (append_3 ([]::int list))"
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
    55
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    56
subsection {* Tricky cases with tuples *}
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    57
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    58
inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    59
where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    60
  "tupled_append ([], xs, xs)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    61
| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    62
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    63
code_pred tupled_append .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    64
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    65
thm tupled_append.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    66
(*
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    67
TODO: values with tupled modes
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    68
values "{xs. tupled_append ([1,2,3], [4,5], xs)}"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
    69
*)
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    70
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    71
inductive tupled_append'
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    72
where
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    73
"tupled_append' ([], xs, xs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    74
| "[| ys = fst (xa, y); x # zs = snd (xa, y);
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    75
 tupled_append' (xs, ys, x # zs) |] ==> tupled_append' (x # xs, xa, y)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    76
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
    77
code_pred (inductify_all) tupled_append' .
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    78
thm tupled_append'.equation
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    79
(* TODO: Modeanalysis returns mode [2] ?? *)
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    80
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    81
inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    82
where
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    83
  "tupled_append'' ([], xs, xs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    84
| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, x # zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    85
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
    86
thm tupled_append''.cases
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
    87
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
    88
code_pred (inductify_all) tupled_append'' .
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    89
thm tupled_append''.equation
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
    90
(* TODO: Modeanalysis returns mode [2] ?? *)
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    91
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    92
inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    93
where
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    94
  "tupled_append''' ([], xs, xs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    95
| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    96
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
    97
code_pred (inductify_all) tupled_append''' .
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    98
thm tupled_append'''.equation
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
    99
(* TODO: Missing a few modes *)
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   100
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   101
inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   102
where
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   103
  "map_ofP ((a, b)#xs) a b"
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   104
| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   105
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   106
code_pred map_ofP .
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   107
thm map_ofP.equation
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   108
section {* reverse *}
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   109
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   110
inductive rev where
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   111
    "rev [] []"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   112
  | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   113
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   114
code_pred rev .
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   115
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   116
thm rev.equation
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   117
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   118
values "{xs. rev [0, 1, 2, 3::nat] xs}"
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   119
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   120
inductive tupled_rev where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   121
  "tupled_rev ([], [])"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   122
| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   123
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   124
code_pred tupled_rev .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   125
thm tupled_rev.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   126
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   127
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   128
  for f where
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   129
    "partition f [] [] []"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   130
  | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   131
  | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   132
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31514
diff changeset
   133
code_pred partition .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   134
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   135
inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   136
  for f where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   137
   "tupled_partition f ([], [], [])"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   138
  | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   139
  | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
32314
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   140
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   141
code_pred tupled_partition .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   142
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   143
thm tupled_partition.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   144
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   145
inductive member
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   146
for xs
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   147
where "x \<in> set xs ==> member xs x"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   148
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   149
lemma [code_pred_intros]:
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   150
  "member (x#xs') x"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   151
by (auto intro: member.intros)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   152
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   153
lemma [code_pred_intros]:
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   154
"member xs x ==> member (x'#xs) x"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   155
by (auto intro: member.intros elim!: member.cases)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   156
(* strange bug must be repaired! *)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   157
(*
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   158
code_pred member sorry
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   159
*)
32314
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   160
inductive is_even :: "nat \<Rightarrow> bool"
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   161
where
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   162
  "n mod 2 = 0 \<Longrightarrow> is_even n"
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   163
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   164
code_pred is_even .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   165
32314
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   166
values 10 "{(ys, zs). partition is_even
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   167
  [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   168
values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   169
values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   170
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31514
diff changeset
   171
lemma [code_pred_intros]:
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   172
  "r a b \<Longrightarrow> tranclp r a b"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   173
  "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   174
  by auto
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   175
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   176
code_pred tranclp
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   177
proof -
31580
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31575
diff changeset
   178
  case tranclp
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31575
diff changeset
   179
  from this converse_tranclpE[OF this(1)] show thesis by metis
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   180
qed
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   181
(*
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   182
code_pred (inductify_all) (rpred) tranclp .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   183
thm tranclp.equation
32316
1d83ac469459 added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents: 32314
diff changeset
   184
thm tranclp.rpred_equation
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   185
*)
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   186
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   187
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   188
    "succ 0 1"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   189
  | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   190
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31514
diff changeset
   191
code_pred succ .
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31514
diff changeset
   192
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   193
thm succ.equation
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   194
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   195
values 10 "{(m, n). succ n m}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   196
values "{m. succ 0 m}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   197
values "{m. succ m 0}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   198
32314
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   199
(* FIXME: why does this not terminate? -- value chooses mode [] --> [1] and then starts enumerating all successors *)
32355
806d2df4d79d properly merged
haftmann
parents: 32351
diff changeset
   200
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   201
(*
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   202
values 20 "{n. tranclp succ 10 n}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   203
values "{n. tranclp succ n 10}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   204
values 20 "{(n, m). tranclp succ n m}"
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   205
*)
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   206
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   207
subsection{* *}
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   208
32424
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   209
subsection{* IMP *}
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   210
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   211
types
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   212
  var = nat
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   213
  state = "int list"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   214
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   215
datatype com =
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   216
  Skip |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   217
  Ass var "state => int" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   218
  Seq com com |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   219
  IF "state => bool" com com |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   220
  While "state => bool" com
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   221
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   222
inductive exec :: "com => state => state => bool" where
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   223
"exec Skip s s" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   224
"exec (Ass x e) s (s[x := e(s)])" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   225
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   226
"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   227
"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   228
"~b s ==> exec (While b c) s s" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   229
"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   230
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   231
code_pred exec .
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   232
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   233
values "{t. exec
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   234
 (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   235
 [3,5] t}"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   236
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   237
inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   238
"tupled_exec (Skip, s, s)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   239
"tupled_exec (Ass x e, s, s[x := e(s)])" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   240
"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   241
"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   242
"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   243
"~b s ==> tupled_exec (While b c, s, s)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   244
"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   245
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   246
code_pred tupled_exec .
32424
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   247
32408
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   248
subsection{* CCS *}
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   249
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   250
text{* This example formalizes finite CCS processes without communication or
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   251
recursion. For simplicity, labels are natural numbers. *}
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   252
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   253
datatype proc = nil | pre nat proc | or proc proc | par proc proc
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   254
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   255
inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   256
"step (pre n p) n p" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   257
"step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   258
"step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   259
"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   260
"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   261
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   262
code_pred step .
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   263
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   264
inductive steps where
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   265
"steps p [] p" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   266
"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   267
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   268
code_pred steps .
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   269
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   270
values 5
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   271
 "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   272
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   273
(* FIXME
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   274
values 3 "{(a,q). step (par nil nil) a q}"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   275
*)
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   276
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   277
inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   278
where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   279
"tupled_step (pre n p, n, p)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   280
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   281
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   282
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   283
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   284
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   285
code_pred tupled_step .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   286
thm tupled_step.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   287
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   288
subsection {* divmod *}
32579
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   289
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   290
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   291
    "k < l \<Longrightarrow> divmod_rel k l 0 k"
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   292
  | "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r"
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   293
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   294
code_pred divmod_rel ..
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   295
33111
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
   296
value [code] "Predicate.the (divmod_rel_1_2 1705 42)"
32579
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   297
32670
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   298
section {* Executing definitions *}
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   299
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   300
definition Min
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   301
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   302
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   303
code_pred (inductify_all) Min .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   304
32670
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   305
subsection {* Examples with lists *}
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   306
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   307
inductive filterP for Pa where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   308
"(filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) [] []"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   309
| "[| (res::'a list) = (y::'a) # (resa::'a list); (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) resa; Pa y |]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   310
==> filterP Pa (y # xt) res"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   311
| "[| (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) (res::'a list); ~ Pa (y::'a) |] ==> filterP Pa (y # xt) res"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   312
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   313
(*
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   314
code_pred (inductify_all) (rpred) filterP .
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   315
thm filterP.rpred_equation
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   316
*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   317
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   318
code_pred (inductify_all) lexord .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   319
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   320
thm lexord.equation
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   321
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   322
lemma "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   323
(*quickcheck[generator=pred_compile]*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   324
oops
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   325
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   326
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   327
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   328
code_pred (inductify_all) lexn .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   329
thm lexn.equation
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   330
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   331
code_pred (inductify_all) lenlex .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   332
thm lenlex.equation
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   333
(*
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   334
code_pred (inductify_all) (rpred) lenlex .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   335
thm lenlex.rpred_equation
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   336
*)
32670
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   337
thm lists.intros
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   338
code_pred (inductify_all) lists .
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   339
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   340
thm lists.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   341
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   342
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   343
fun height :: "'a tree => nat" where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   344
"height ET = 0"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   345
| "height (MKT x l r h) = max (height l) (height r) + 1"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   346
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   347
consts avl :: "'a tree => bool"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   348
primrec
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   349
  "avl ET = True"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   350
  "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and> 
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   351
  h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   352
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   353
code_pred (inductify_all) avl .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   354
thm avl.equation
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   355
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   356
lemma [code_pred_inline]: "bot_fun_inst.bot_fun == (\<lambda>(y::'a::type). False)"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   357
unfolding bot_fun_inst.bot_fun[symmetric] bot_bool_eq[symmetric] bot_fun_eq by simp
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   358
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   359
fun set_of
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   360
where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   361
"set_of ET = {}"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   362
| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   363
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   364
fun is_ord
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   365
where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   366
"is_ord ET = True"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   367
| "is_ord (MKT n l r h) =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   368
 ((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   369
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   370
declare Un_def[code_pred_def]
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   371
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   372
code_pred (inductify_all) set_of .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   373
thm set_of.equation
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   374
(* FIXME *)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   375
(*
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   376
code_pred (inductify_all) is_ord .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   377
thm is_ord.equation
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   378
*)
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   379
section {* Definitions about Relations *}
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   380
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   381
code_pred (inductify_all) converse .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   382
thm converse.equation
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   383
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   384
code_pred (inductify_all) Domain .
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   385
thm Domain.equation
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   386
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   387
section {* List functions *}
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   388
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   389
code_pred (inductify_all) length .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   390
thm size_listP.equation
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   391
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   392
code_pred (inductify_all) concat .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   393
thm concatP.equation
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   394
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   395
code_pred (inductify_all) hd .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   396
code_pred (inductify_all) tl .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   397
code_pred (inductify_all) last .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   398
code_pred (inductify_all) butlast .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   399
(*code_pred (inductify_all) listsum .*)
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   400
code_pred (inductify_all) take .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   401
code_pred (inductify_all) drop .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   402
code_pred (inductify_all) zip .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   403
code_pred (inductify_all) upt .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   404
code_pred set sorry
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   405
code_pred (inductify_all) remdups .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   406
code_pred (inductify_all) remove1 .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   407
code_pred (inductify_all) removeAll .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   408
code_pred (inductify_all) distinct .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   409
code_pred (inductify_all) replicate .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   410
code_pred (inductify_all) splice .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   411
code_pred (inductify_all) List.rev .
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   412
thm revP.equation
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   413
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   414
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   415
section {* Context Free Grammar *}
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   416
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   417
datatype alphabet = a | b
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   418
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   419
inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   420
  "[] \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   421
| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   422
| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   423
| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   424
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   425
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   426
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   427
code_pred (inductify_all) S\<^isub>1p .
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   428
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   429
thm S\<^isub>1p.equation
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33105
diff changeset
   430
(*
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   431
theorem S\<^isub>1_sound:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   432
"w \<in> S\<^isub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   433
quickcheck[generator=pred_compile]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   434
oops
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   435
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   436
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   437
  "[] \<in> S\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   438
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   439
| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   440
| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   441
| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   442
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   443
(*
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   444
code_pred (inductify_all) (rpred) S\<^isub>2 .
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   445
ML {* Predicate_Compile_Core.intros_of @{theory} @{const_name "B\<^isub>2"} *}
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   446
*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   447
theorem S\<^isub>2_sound:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   448
"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   449
(*quickcheck[generator=SML]*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   450
quickcheck[generator=pred_compile, size=15, iterations=100]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   451
oops
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   452
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   453
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   454
  "[] \<in> S\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   455
| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   456
| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   457
| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   458
| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   459
| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   460
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   461
(*
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   462
code_pred (inductify_all) S\<^isub>3 .
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   463
*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   464
theorem S\<^isub>3_sound:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   465
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   466
quickcheck[generator=pred_compile, size=10, iterations=1]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   467
oops
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   468
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   469
lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   470
quickcheck[size=10, generator = pred_compile]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   471
oops
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33105
diff changeset
   472
*)
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33105
diff changeset
   473
inductive test
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33105
diff changeset
   474
where
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33105
diff changeset
   475
  "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] ==> test w"
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33105
diff changeset
   476
ML {* @{term "[x \<leftarrow> w. x = a]"} *}
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33105
diff changeset
   477
code_pred (inductify_all) test .
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   478
(*
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   479
theorem S\<^isub>3_complete:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   480
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   481
(*quickcheck[generator=SML]*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   482
quickcheck[generator=pred_compile, size=10, iterations=100]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   483
oops
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   484
*)
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   485
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   486
  "[] \<in> S\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   487
| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   488
| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   489
| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   490
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   491
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   492
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   493
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   494
theorem S\<^isub>4_sound:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   495
"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   496
quickcheck[generator = pred_compile, size=2, iterations=1]
32673
d5db9cf85401 replaced sorry by oops; removed old debug functions in predicate compiler
bulwahn
parents: 32672
diff changeset
   497
oops
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   498
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   499
theorem S\<^isub>4_complete:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   500
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   501
quickcheck[generator = pred_compile, size=5, iterations=1]
32673
d5db9cf85401 replaced sorry by oops; removed old debug functions in predicate compiler
bulwahn
parents: 32672
diff changeset
   502
oops
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   503
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   504
theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   505
"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   506
"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   507
"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   508
(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
32673
d5db9cf85401 replaced sorry by oops; removed old debug functions in predicate compiler
bulwahn
parents: 32672
diff changeset
   509
oops
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   510
33104
560372b461e5 moved meta_fun_cong lemma into ML-file; tuned
bulwahn
parents: 32673
diff changeset
   511
section {* Lambda *}
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   512
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   513
datatype type =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   514
    Atom nat
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   515
  | Fun type type    (infixr "\<Rightarrow>" 200)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   516
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   517
datatype dB =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   518
    Var nat
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   519
  | App dB dB (infixl "\<degree>" 200)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   520
  | Abs type dB
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   521
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   522
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   523
  nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   524
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   525
  "[]\<langle>i\<rangle> = None"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   526
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   527
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   528
(*
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   529
inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   530
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   531
  "nth_el' (x # xs) 0 x"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   532
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   533
*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   534
inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   535
  where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   536
    Var [intro!]: "nth_el env x = Some T \<Longrightarrow> env \<turnstile> Var x : T"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   537
  | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   538
(*  | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" *)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   539
  | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   540
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   541
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   542
  lift :: "[dB, nat] => dB"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   543
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   544
    "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   545
  | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   546
  | "lift (Abs T s) k = Abs T (lift s (k + 1))"
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   547
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   548
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   549
  subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   550
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   551
    subst_Var: "(Var i)[s/k] =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   552
      (if k < i then Var (i - 1) else if i = k then s else Var i)"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   553
  | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   554
  | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   555
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   556
inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   557
  where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   558
    beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   559
  | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   560
  | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   561
  | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   562
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   563
lemma "Gamma \<turnstile> t : T \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> Gamma \<turnstile> t' : T"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   564
quickcheck[generator = pred_compile, size = 10, iterations = 1000]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   565
oops
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   566
(* FIXME *)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   567
(*
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   568
inductive test for P where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   569
"[| filter P vs = res |]
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   570
==> test P vs res"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   571
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   572
code_pred test .
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   573
*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   574
(*
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   575
export_code test_for_1_yields_1_2 in SML file -
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   576
code_pred (inductify_all) (rpred) test .
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   577
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   578
thm test.equation
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   579
*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   580
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   581
lemma filter_eq_ConsD:
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   582
 "filter P ys = x#xs \<Longrightarrow>
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   583
  \<exists>us vs. ys = ts @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   584
(*quickcheck[generator = pred_compile]*)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   585
oops
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   586
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   587
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   588
end