| author | haftmann |
| Wed, 28 Oct 2009 19:09:47 +0100 | |
| changeset 33296 | a3924d1069e5 |
| parent 33258 | f47ca46d2187 |
| child 33326 | 7d0288d90535 |
| permissions | -rw-r--r-- |
|
31129
d2cead76fca2
split Predicate_Compile examples into separate theory
haftmann
parents:
31123
diff
changeset
|
1 |
theory Predicate_Compile_ex |
| 33120 | 2 |
imports Main Predicate_Compile_Alternative_Defs |
|
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 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
5 |
subsection {* Basic predicates *}
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
6 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
7 |
inductive False' :: "bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
8 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
9 |
code_pred (mode: []) False' . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
10 |
code_pred [depth_limited] False' . |
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
11 |
code_pred [rpred] False' . |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
12 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
13 |
inductive EmptySet :: "'a \<Rightarrow> bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
14 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
15 |
code_pred (mode: [], [1]) EmptySet . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
16 |
|
| 33253 | 17 |
definition EmptySet' :: "'a \<Rightarrow> bool" |
18 |
where "EmptySet' = {}"
|
|
19 |
||
20 |
code_pred (mode: [], [1]) [inductify, show_intermediate_results] EmptySet' . |
|
21 |
||
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
22 |
inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
23 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
24 |
code_pred (mode: [], [1], [2], [1, 2]) EmptyRel . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
25 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
26 |
inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
27 |
for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
28 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
29 |
code_pred (mode: [], [1], [2], [1, 2])EmptyClosure . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
30 |
thm EmptyClosure.equation |
|
33258
f47ca46d2187
disabled test case for predicate compiler due to an problem in the inductive package
bulwahn
parents:
33255
diff
changeset
|
31 |
(* TODO: inductive package is broken! |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
32 |
inductive False'' :: "bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
33 |
where |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
34 |
"False \<Longrightarrow> False''" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
35 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
36 |
code_pred (mode: []) False'' . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
37 |
|
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
38 |
inductive EmptySet'' :: "'a \<Rightarrow> bool" |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
39 |
where |
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
40 |
"False \<Longrightarrow> EmptySet'' x" |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
41 |
|
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
42 |
code_pred (mode: [1]) EmptySet'' . |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
43 |
code_pred (mode: [], [1]) [inductify] EmptySet'' . |
|
33258
f47ca46d2187
disabled test case for predicate compiler due to an problem in the inductive package
bulwahn
parents:
33255
diff
changeset
|
44 |
*) |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
45 |
inductive True' :: "bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
46 |
where |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
47 |
"True \<Longrightarrow> True'" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
48 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
49 |
code_pred (mode: []) True' . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
50 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
51 |
consts a' :: 'a |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
52 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
53 |
inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
54 |
where |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
55 |
"Fact a' a'" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
56 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
57 |
code_pred (mode: [], [1], [2], [1, 2]) Fact . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
58 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
59 |
inductive zerozero :: "nat * nat => bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
60 |
where |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
61 |
"zerozero (0, 0)" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
62 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
63 |
code_pred zerozero . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
64 |
code_pred [rpred, show_compilation] zerozero . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
65 |
|
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
66 |
subsection {* Preprocessor Inlining *}
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
67 |
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
68 |
definition "equals == (op =)" |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
69 |
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
70 |
inductive zerozero' where |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
71 |
"equals (x, y) (0, 0) ==> zerozero' (x, y)" |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
72 |
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
73 |
code_pred (mode: [1]) zerozero' . |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
74 |
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
75 |
lemma zerozero'_eq: "zerozero' == zerozero" |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
76 |
proof - |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
77 |
have "zerozero' = zerozero" |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
78 |
apply (auto simp add: mem_def) |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
79 |
apply (cases rule: zerozero'.cases) |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
80 |
apply (auto simp add: equals_def intro: zerozero.intros) |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
81 |
apply (cases rule: zerozero.cases) |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
82 |
apply (auto simp add: equals_def intro: zerozero'.intros) |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
83 |
done |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
84 |
from this show "zerozero' == zerozero" by auto |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
85 |
qed |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
86 |
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
87 |
declare zerozero'_eq [code_pred_inline] |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
88 |
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
89 |
definition "zerozero'' x == zerozero' x" |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
90 |
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
91 |
text {* if preprocessing fails, zerozero'' will not have all modes. *}
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
92 |
ML {* Predicate_Compile_Inline_Defs.get @{context} *}
|
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
93 |
(* TODO: *) |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
94 |
code_pred (mode: [1]) [inductify, show_intermediate_results] zerozero'' . |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
95 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
96 |
subsection {* even predicate *}
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
97 |
|
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
98 |
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
|
99 |
"even 0" |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
100 |
| "even n \<Longrightarrow> odd (Suc n)" |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
101 |
| "odd n \<Longrightarrow> even (Suc n)" |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
102 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
103 |
code_pred (mode: [], [1]) even . |
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
104 |
code_pred [depth_limited] even . |
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
105 |
code_pred [rpred] even . |
|
33137
0d16c07f8d24
added option to generate random values to values command in the predicate compiler
bulwahn
parents:
33136
diff
changeset
|
106 |
|
| 31514 | 107 |
thm odd.equation |
| 31123 | 108 |
thm even.equation |
|
33136
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
109 |
thm odd.depth_limited_equation |
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
110 |
thm even.depth_limited_equation |
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
111 |
thm even.rpred_equation |
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
112 |
thm odd.rpred_equation |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
113 |
|
| 31217 | 114 |
values "{x. even 2}"
|
115 |
values "{x. odd 2}"
|
|
116 |
values 10 "{n. even n}"
|
|
117 |
values 10 "{n. odd n}"
|
|
|
33136
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
118 |
values [depth_limit = 2] "{x. even 6}"
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
119 |
values [depth_limit = 7] "{x. even 6}"
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
120 |
values [depth_limit = 2] "{x. odd 7}"
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
121 |
values [depth_limit = 8] "{x. odd 7}"
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
122 |
values [depth_limit = 7] 10 "{n. even n}"
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
123 |
|
|
33136
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
124 |
definition odd' where "odd' x == \<not> even x" |
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
125 |
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
126 |
code_pred [inductify] odd' . |
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
127 |
code_pred [inductify, depth_limited] odd' . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
128 |
code_pred [inductify, rpred] odd' . |
|
33136
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
129 |
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
130 |
thm odd'.depth_limited_equation |
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
131 |
values [depth_limit = 2] "{x. odd' 7}"
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
132 |
values [depth_limit = 9] "{x. odd' 7}"
|
|
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
133 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
134 |
inductive is_even :: "nat \<Rightarrow> bool" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
135 |
where |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
136 |
"n mod 2 = 0 \<Longrightarrow> is_even n" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
137 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
138 |
code_pred is_even . |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
139 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
140 |
subsection {* append predicate *}
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
141 |
|
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
142 |
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where |
| 32340 | 143 |
"append [] xs xs" |
144 |
| "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
|
145 |
|
|
33146
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
146 |
code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append . |
|
33136
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
147 |
code_pred [depth_limited] append . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
148 |
code_pred [rpred] append . |
| 31123 | 149 |
|
150 |
thm append.equation |
|
|
33136
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
151 |
thm append.depth_limited_equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
152 |
thm append.rpred_equation |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
153 |
|
| 31217 | 154 |
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
|
155 |
values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
|
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
156 |
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
|
|
33136
74d51fb3be2e
commented out the random generator compilation in the example file
bulwahn
parents:
33129
diff
changeset
|
157 |
values [depth_limit = 3] "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
|
|
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
158 |
values [random] 15 "{(ys, zs). append [1::nat, 2] ys zs}"
|
| 31195 | 159 |
|
| 33111 | 160 |
value [code] "Predicate.the (append_1_2 [0::int, 1, 2] [3, 4, 5])" |
161 |
value [code] "Predicate.the (append_3 ([]::int list))" |
|
162 |
||
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
163 |
text {* tricky case with alternative rules *}
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
164 |
|
|
33116
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
165 |
inductive append2 |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
166 |
where |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
167 |
"append2 [] xs xs" |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
168 |
| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)" |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
169 |
|
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
170 |
lemma append2_Nil: "append2 [] (xs::'b list) xs" |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
171 |
by (simp add: append2.intros(1)) |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
172 |
|
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
173 |
lemmas [code_pred_intros] = append2_Nil append2.intros(2) |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
174 |
|
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
175 |
code_pred append2 |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
176 |
proof - |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
177 |
case append2 |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
178 |
from append2.cases[OF append2(1)] append2(2-3) show thesis by blast |
|
b379ee2cddb1
changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents:
33114
diff
changeset
|
179 |
qed |
| 33128 | 180 |
|
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
181 |
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
|
182 |
where |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
183 |
"tupled_append ([], xs, xs)" |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
184 |
| "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
|
185 |
|
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
186 |
code_pred tupled_append . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
187 |
code_pred [rpred] tupled_append . |
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
188 |
thm tupled_append.equation |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
189 |
(* |
|
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
|
190 |
TODO: values with tupled modes |
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
191 |
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
|
192 |
*) |
|
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
|
193 |
|
|
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
|
194 |
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
|
195 |
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
|
196 |
"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
|
197 |
| "[| ys = fst (xa, y); x # zs = snd (xa, y); |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
198 |
tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)" |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
199 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
200 |
code_pred 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
|
201 |
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
|
202 |
|
|
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
|
203 |
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
|
204 |
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
|
205 |
"tupled_append'' ([], xs, xs)" |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
206 |
| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)" |
|
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 |
|
| 33113 | 208 |
thm tupled_append''.cases |
209 |
||
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
210 |
code_pred [inductify] 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
|
211 |
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
|
212 |
|
|
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
|
213 |
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
|
214 |
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
|
215 |
"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
|
216 |
| "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
|
217 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
218 |
code_pred [inductify] 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
|
219 |
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
|
220 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
221 |
subsection {* map_ofP predicate *}
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
222 |
|
| 33113 | 223 |
inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
|
224 |
where |
|
225 |
"map_ofP ((a, b)#xs) a b" |
|
226 |
| "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
|
227 |
|
| 33120 | 228 |
code_pred (mode: [1], [1, 2], [1, 2, 3], [1, 3]) map_ofP . |
| 33113 | 229 |
thm map_ofP.equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
230 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
231 |
subsection {* filter predicate *}
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
232 |
|
|
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
233 |
inductive filter1 |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
234 |
for P |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
235 |
where |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
236 |
"filter1 P [] []" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
237 |
| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
238 |
| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
239 |
|
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
240 |
code_pred (mode: [1], [1, 2]) filter1 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
241 |
code_pred [depth_limited] filter1 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
242 |
code_pred [rpred] filter1 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
243 |
|
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
244 |
thm filter1.equation |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
245 |
|
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
246 |
inductive filter2 |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
247 |
where |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
248 |
"filter2 P [] []" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
249 |
| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
250 |
| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
251 |
|
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
252 |
code_pred (mode: [1, 2, 3], [1, 2]) filter2 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
253 |
code_pred [depth_limited] filter2 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
254 |
code_pred [rpred] filter2 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
255 |
thm filter2.equation |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
256 |
thm filter2.rpred_equation |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
257 |
|
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
258 |
inductive filter3 |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
259 |
for P |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
260 |
where |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
261 |
"List.filter P xs = ys ==> filter3 P xs ys" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
262 |
|
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
263 |
code_pred filter3 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
264 |
code_pred [depth_limited] filter3 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
265 |
thm filter3.depth_limited_equation |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
266 |
|
|
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
267 |
inductive filter4 |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
268 |
where |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
269 |
"List.filter P xs = ys ==> filter4 P xs ys" |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
270 |
|
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
271 |
code_pred filter4 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
272 |
code_pred [depth_limited] filter4 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
273 |
code_pred [rpred] filter4 . |
|
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
274 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
275 |
subsection {* reverse predicate *}
|
|
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
|
276 |
|
| 32340 | 277 |
inductive rev where |
278 |
"rev [] []" |
|
279 |
| "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys" |
|
280 |
||
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
281 |
code_pred (mode: [1], [2], [1, 2]) rev . |
| 32340 | 282 |
|
283 |
thm rev.equation |
|
284 |
||
285 |
values "{xs. rev [0, 1, 2, 3::nat] xs}"
|
|
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
286 |
|
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
287 |
inductive tupled_rev where |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
288 |
"tupled_rev ([], [])" |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
289 |
| "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
|
290 |
|
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
291 |
code_pred tupled_rev . |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
292 |
thm tupled_rev.equation |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
293 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
294 |
subsection {* partition predicate *}
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
295 |
|
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
296 |
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
|
297 |
for f where |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
298 |
"partition f [] [] []" |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
299 |
| "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
|
300 |
| "\<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
|
301 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
302 |
code_pred (mode: [1], [2, 3], [1, 2], [1, 3], [1, 2, 3]) partition . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
303 |
code_pred [depth_limited] partition . |
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
304 |
code_pred [rpred] partition . |
| 31123 | 305 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
306 |
values 10 "{(ys, zs). partition is_even
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
307 |
[0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
308 |
values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
309 |
values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
310 |
|
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
311 |
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
|
312 |
for f where |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
313 |
"tupled_partition f ([], [], [])" |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
314 |
| "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
|
315 |
| "\<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
|
316 |
|
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
317 |
code_pred tupled_partition . |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
318 |
|
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
319 |
thm tupled_partition.equation |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
320 |
|
|
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
|
321 |
lemma [code_pred_intros]: |
| 32340 | 322 |
"r a b \<Longrightarrow> tranclp r a b" |
323 |
"r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c" |
|
324 |
by auto |
|
|
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
325 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
326 |
subsection {* transitive predicate *}
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
327 |
|
|
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
328 |
code_pred tranclp |
|
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
329 |
proof - |
| 31580 | 330 |
case tranclp |
331 |
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
|
332 |
qed |
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
333 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
334 |
code_pred [depth_limited] tranclp . |
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
335 |
code_pred [rpred] tranclp . |
| 31123 | 336 |
thm tranclp.equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
337 |
thm tranclp.rpred_equation |
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
338 |
|
| 31217 | 339 |
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where |
340 |
"succ 0 1" |
|
341 |
| "succ m n \<Longrightarrow> succ (Suc m) (Suc n)" |
|
342 |
||
|
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
|
343 |
code_pred succ . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
344 |
code_pred [rpred] succ . |
| 31217 | 345 |
thm succ.equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
346 |
thm succ.rpred_equation |
| 32340 | 347 |
|
348 |
values 10 "{(m, n). succ n m}"
|
|
349 |
values "{m. succ 0 m}"
|
|
350 |
values "{m. succ m 0}"
|
|
351 |
||
|
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
352 |
(* FIXME: why does this not terminate? -- value chooses mode [] --> [1] and then starts enumerating all successors *) |
| 32355 | 353 |
|
| 31514 | 354 |
(* |
| 31217 | 355 |
values 20 "{n. tranclp succ 10 n}"
|
356 |
values "{n. tranclp succ n 10}"
|
|
357 |
values 20 "{(n, m). tranclp succ n m}"
|
|
| 31514 | 358 |
*) |
| 31217 | 359 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
360 |
subsection {* IMP *}
|
| 32424 | 361 |
|
362 |
types |
|
363 |
var = nat |
|
364 |
state = "int list" |
|
365 |
||
366 |
datatype com = |
|
367 |
Skip | |
|
368 |
Ass var "state => int" | |
|
369 |
Seq com com | |
|
370 |
IF "state => bool" com com | |
|
371 |
While "state => bool" com |
|
372 |
||
373 |
inductive exec :: "com => state => state => bool" where |
|
374 |
"exec Skip s s" | |
|
375 |
"exec (Ass x e) s (s[x := e(s)])" | |
|
376 |
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" | |
|
377 |
"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" | |
|
378 |
"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" | |
|
379 |
"~b s ==> exec (While b c) s s" | |
|
380 |
"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3" |
|
381 |
||
382 |
code_pred exec . |
|
383 |
||
384 |
values "{t. exec
|
|
385 |
(While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1)))) |
|
386 |
[3,5] t}" |
|
387 |
||
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
388 |
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
|
389 |
"tupled_exec (Skip, s, s)" | |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
390 |
"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
|
391 |
"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
|
392 |
"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
|
393 |
"~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
|
394 |
"~b s ==> tupled_exec (While b c, s, s)" | |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
395 |
"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
|
396 |
|
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
397 |
code_pred tupled_exec . |
| 32424 | 398 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
399 |
subsection {* CCS *}
|
| 32408 | 400 |
|
401 |
text{* This example formalizes finite CCS processes without communication or
|
|
402 |
recursion. For simplicity, labels are natural numbers. *} |
|
403 |
||
404 |
datatype proc = nil | pre nat proc | or proc proc | par proc proc |
|
405 |
||
406 |
inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where |
|
407 |
"step (pre n p) n p" | |
|
408 |
"step p1 a q \<Longrightarrow> step (or p1 p2) a q" | |
|
409 |
"step p2 a q \<Longrightarrow> step (or p1 p2) a q" | |
|
410 |
"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" | |
|
411 |
"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)" |
|
412 |
||
413 |
code_pred step . |
|
414 |
||
415 |
inductive steps where |
|
416 |
"steps p [] p" | |
|
417 |
"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r" |
|
418 |
||
419 |
code_pred steps . |
|
420 |
||
421 |
values 5 |
|
422 |
"{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
|
|
423 |
||
424 |
(* FIXME |
|
425 |
values 3 "{(a,q). step (par nil nil) a q}"
|
|
426 |
*) |
|
427 |
||
|
33105
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
428 |
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
|
429 |
where |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
430 |
"tupled_step (pre n p, n, p)" | |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
431 |
"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
|
432 |
"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
|
433 |
"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
|
434 |
"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
|
435 |
|
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
436 |
code_pred tupled_step . |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
437 |
thm tupled_step.equation |
|
1e4146bf841c
added tupled versions of examples for the predicate compiler
bulwahn
parents:
33104
diff
changeset
|
438 |
|
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
439 |
subsection {* divmod *}
|
| 32579 | 440 |
|
441 |
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where |
|
442 |
"k < l \<Longrightarrow> divmod_rel k l 0 k" |
|
443 |
| "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r" |
|
444 |
||
445 |
code_pred divmod_rel .. |
|
446 |
||
| 33111 | 447 |
value [code] "Predicate.the (divmod_rel_1_2 1705 42)" |
| 32579 | 448 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
449 |
subsection {* Minimum *}
|
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
450 |
|
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
451 |
definition Min |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
452 |
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
|
453 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
454 |
code_pred [inductify] Min . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
455 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
456 |
subsection {* Lexicographic order *}
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
457 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
458 |
code_pred [inductify] lexord . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
459 |
code_pred [inductify, rpred] lexord . |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
460 |
thm lexord.equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
461 |
thm lexord.rpred_equation |
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
462 |
|
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
463 |
inductive less_than_nat :: "nat * nat => bool" |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
464 |
where |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
465 |
"less_than_nat (0, x)" |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
466 |
| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)" |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
467 |
|
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
468 |
code_pred less_than_nat . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
469 |
|
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
470 |
code_pred [depth_limited] less_than_nat . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
471 |
code_pred [rpred] less_than_nat . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
472 |
|
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
473 |
inductive test_lexord :: "nat list * nat list => bool" |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
474 |
where |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
475 |
"lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)" |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
476 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
477 |
code_pred [rpred] test_lexord . |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
478 |
code_pred [depth_limited] test_lexord . |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
479 |
thm test_lexord.depth_limited_equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
480 |
thm test_lexord.rpred_equation |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
481 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
482 |
values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
|
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
483 |
values [depth_limit = 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
484 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
485 |
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
486 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
487 |
code_pred [inductify] lexn . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
488 |
thm lexn.equation |
|
33139
9c01ee6f8ee9
added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents:
33138
diff
changeset
|
489 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
490 |
code_pred [rpred] lexn . |
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
491 |
|
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
492 |
thm lexn.rpred_equation |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
493 |
|
|
33146
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
494 |
code_pred [inductify, show_steps] lenlex . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
495 |
thm lenlex.equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
496 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
497 |
code_pred [inductify, rpred] lenlex . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
498 |
thm lenlex.rpred_equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
499 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
500 |
code_pred [inductify] lists . |
| 32670 | 501 |
thm lists.equation |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
502 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
503 |
subsection {* AVL Tree *}
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
504 |
|
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
505 |
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
|
506 |
fun height :: "'a tree => nat" where |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
507 |
"height ET = 0" |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
508 |
| "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
|
509 |
|
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
510 |
consts avl :: "'a tree => bool" |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
511 |
primrec |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
512 |
"avl ET = True" |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
513 |
"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
|
514 |
h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)" |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
515 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
516 |
code_pred [inductify] avl . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
517 |
thm avl.equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
518 |
|
|
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
519 |
code_pred [rpred] avl . |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
520 |
thm avl.rpred_equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
521 |
|
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
522 |
fun set_of |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
523 |
where |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
524 |
"set_of ET = {}"
|
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
525 |
| "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
|
526 |
|
|
33123
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
bulwahn
parents:
33120
diff
changeset
|
527 |
fun is_ord :: "nat tree => bool" |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
528 |
where |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
529 |
"is_ord ET = True" |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
530 |
| "is_ord (MKT n l r h) = |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
531 |
((\<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
|
532 |
|
|
33140
83951822bfd0
cleaning the signature of the predicate compiler core; renaming signature and structures to uniform long names
bulwahn
parents:
33139
diff
changeset
|
533 |
code_pred (mode: [1], [1, 2]) [inductify] set_of . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
534 |
thm set_of.equation |
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
535 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
536 |
code_pred [inductify] is_ord . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
537 |
thm is_ord.equation |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
538 |
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
539 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
540 |
subsection {* Definitions about Relations *}
|
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
541 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
542 |
code_pred [inductify] converse . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
543 |
thm converse.equation |
| 33145 | 544 |
code_pred [inductify] rel_comp . |
545 |
thm rel_comp.equation |
|
546 |
code_pred [inductify] Image . |
|
547 |
thm Image.equation |
|
548 |
(*TODO: *) |
|
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
549 |
|
|
33146
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
550 |
declare Id_on_def[unfolded UNION_def, code_pred_def] |
|
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
551 |
|
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
552 |
code_pred [inductify] Id_on . |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
553 |
thm Id_on.equation |
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
554 |
code_pred [inductify] Domain . |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
555 |
thm Domain.equation |
| 33145 | 556 |
code_pred [inductify] Range . |
| 33253 | 557 |
thm Range.equation |
| 33145 | 558 |
code_pred [inductify] Field . |
|
33146
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
559 |
declare Sigma_def[unfolded UNION_def, code_pred_def] |
|
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
560 |
declare refl_on_def[unfolded UNION_def, code_pred_def] |
|
33255
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
561 |
code_pred [inductify] refl_on . |
|
75b01355e5d4
adding test case for inlining of predicate compiler
bulwahn
parents:
33253
diff
changeset
|
562 |
thm refl_on.equation |
| 33145 | 563 |
code_pred [inductify] total_on . |
564 |
thm total_on.equation |
|
565 |
code_pred [inductify] antisym . |
|
566 |
thm antisym.equation |
|
567 |
code_pred [inductify] trans . |
|
568 |
thm trans.equation |
|
569 |
code_pred [inductify] single_valued . |
|
570 |
thm single_valued.equation |
|
571 |
code_pred [inductify] inv_image . |
|
572 |
thm inv_image.equation |
|
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
573 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
574 |
subsection {* Inverting list functions *}
|
|
33114
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents:
33113
diff
changeset
|
575 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
576 |
code_pred [inductify] length . |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
577 |
code_pred [inductify, rpred] length . |
|
33114
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents:
33113
diff
changeset
|
578 |
thm size_listP.equation |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
579 |
thm size_listP.rpred_equation |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
580 |
|
|
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
581 |
values [random] 20 "{xs. size_listP (xs::nat list) (5::nat)}"
|
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
582 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
583 |
code_pred [inductify] concat . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
584 |
code_pred [inductify] hd . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
585 |
code_pred [inductify] tl . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
586 |
code_pred [inductify] last . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
587 |
code_pred [inductify] butlast . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
588 |
code_pred [inductify] take . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
589 |
code_pred [inductify] drop . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
590 |
code_pred [inductify] zip . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
591 |
code_pred [inductify] upt . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
592 |
code_pred [inductify] remdups . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
593 |
code_pred [inductify] remove1 . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
594 |
code_pred [inductify] removeAll . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
595 |
code_pred [inductify] distinct . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
596 |
code_pred [inductify] replicate . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
597 |
code_pred [inductify] splice . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
598 |
code_pred [inductify] List.rev . |
|
33129
3085da75ed54
changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents:
33128
diff
changeset
|
599 |
code_pred [inductify] map . |
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
600 |
code_pred [inductify] foldr . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
601 |
code_pred [inductify] foldl . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
602 |
code_pred [inductify] filter . |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
603 |
code_pred [inductify, rpred] filter . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
604 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
605 |
subsection {* Context Free Grammar *}
|
|
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
|
606 |
|
|
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
607 |
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
|
608 |
|
|
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
609 |
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
|
610 |
"[] \<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
|
611 |
| "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
|
612 |
| "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
|
613 |
| "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
|
614 |
| "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
|
615 |
| "\<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
|
616 |
|
|
33140
83951822bfd0
cleaning the signature of the predicate compiler core; renaming signature and structures to uniform long names
bulwahn
parents:
33139
diff
changeset
|
617 |
code_pred [inductify] S\<^isub>1p . |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
618 |
code_pred [inductify, rpred] S\<^isub>1p . |
|
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
|
619 |
thm S\<^isub>1p.equation |
|
33141
89b23fad5e02
modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents:
33140
diff
changeset
|
620 |
thm S\<^isub>1p.rpred_equation |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
621 |
|
|
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
622 |
values [random] 5 "{x. S\<^isub>1p x}"
|
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
623 |
|
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
624 |
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
|
625 |
"[] \<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
|
626 |
| "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
|
627 |
| "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
|
628 |
| "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
|
629 |
| "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
|
630 |
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2" |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
631 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
632 |
code_pred [inductify, rpred] S\<^isub>2 . |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
633 |
thm S\<^isub>2.rpred_equation |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
634 |
thm A\<^isub>2.rpred_equation |
|
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
635 |
thm B\<^isub>2.rpred_equation |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
636 |
|
|
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
637 |
values [random] 10 "{x. S\<^isub>2 x}"
|
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
638 |
|
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
639 |
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
|
640 |
"[] \<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
|
641 |
| "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
|
642 |
| "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
|
643 |
| "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
|
644 |
| "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
|
645 |
| "\<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
|
646 |
|
|
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
647 |
code_pred [inductify] S\<^isub>3 . |
|
33138
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
648 |
thm S\<^isub>3.equation |
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
649 |
|
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
650 |
values 10 "{x. S\<^isub>3 x}"
|
|
e2e23987c59a
reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents:
33137
diff
changeset
|
651 |
|
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
652 |
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
|
653 |
"[] \<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
|
654 |
| "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
|
655 |
| "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
|
656 |
| "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
|
657 |
| "\<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
|
658 |
| "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
|
659 |
| "\<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
|
660 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
661 |
subsection {* 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
|
662 |
|
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
663 |
datatype type = |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
664 |
Atom nat |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
665 |
| 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
|
666 |
|
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
667 |
datatype dB = |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
668 |
Var nat |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
669 |
| 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
|
670 |
| Abs type dB |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
671 |
|
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
672 |
primrec |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
673 |
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
|
674 |
where |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
675 |
"[]\<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
|
676 |
| "(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
|
677 |
|
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
678 |
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
|
679 |
where |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
680 |
"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
|
681 |
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y" |
| 33128 | 682 |
|
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
683 |
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
|
684 |
where |
| 33128 | 685 |
Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T" |
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
686 |
| Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)" |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
33147
diff
changeset
|
687 |
| App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" |
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
688 |
|
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
689 |
primrec |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
690 |
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
|
691 |
where |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
692 |
"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
|
693 |
| "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
|
694 |
| "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
|
695 |
|
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
696 |
primrec |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
697 |
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
|
698 |
where |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
699 |
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
|
700 |
(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
|
701 |
| 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
|
702 |
| 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
|
703 |
|
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
704 |
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
|
705 |
where |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
706 |
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
|
707 |
| 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
|
708 |
| 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
|
709 |
| 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
|
710 |
|
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
711 |
end |