| author | paulson <lp15@cam.ac.uk> | 
| Mon, 07 Dec 2015 16:44:26 +0000 | |
| changeset 61806 | d2e62ae01cd8 | 
| parent 60565 | b7ee41f72add | 
| child 62121 | c8a93680b80d | 
| permissions | -rw-r--r-- | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
theory Predicate_Compile_Tests  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40885 
diff
changeset
 | 
2  | 
imports "~~/src/HOL/Library/Predicate_Compile_Alternative_Defs"  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
begin  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
|
| 
42208
 
02513eb26eb7
raised timeouts further, for SML/NJ -- because of variations in machines/compilers, fixed timeouts can merely prevent non-termination, not enforce particular performance characteristics.
 
krauss 
parents: 
42142 
diff
changeset
 | 
5  | 
declare [[values_timeout = 480.0]]  | 
| 
42142
 
d24a93025feb
raised various timeouts to accommodate sluggish SML/NJ
 
krauss 
parents: 
41413 
diff
changeset
 | 
6  | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
subsection {* Basic predicates *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
inductive False' :: "bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
code_pred (expected_modes: bool) False' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
code_pred [dseq] False' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
code_pred [random_dseq] False' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
values [expected "{}" pred] "{x. False'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
values [expected "{}" dseq 1] "{x. False'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
values [expected "{}" random_dseq 1, 1, 1] "{x. False'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
value "False'"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
inductive True' :: "bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
"True ==> True'"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
code_pred True' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
code_pred [dseq] True' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
27  | 
code_pred [random_dseq] True' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
thm True'.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
thm True'.dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
thm True'.random_dseq_equation  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
32  | 
values [expected "{()}"] "{x. True'}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
values [expected "{}" dseq 0] "{x. True'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
values [expected "{()}" dseq 1] "{x. True'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
values [expected "{()}" dseq 2] "{x. True'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
36  | 
values [expected "{}" random_dseq 1, 1, 0] "{x. True'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
values [expected "{}" random_dseq 1, 1, 1] "{x. True'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
values [expected "{()}" random_dseq 1, 1, 2] "{x. True'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
39  | 
values [expected "{()}" random_dseq 1, 1, 3] "{x. True'}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
41  | 
inductive EmptyPred :: "'a \<Rightarrow> bool"  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
43  | 
code_pred (expected_modes: o => bool, i => bool) EmptyPred .  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
45  | 
definition EmptyPred' :: "'a \<Rightarrow> bool"  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
46  | 
where "EmptyPred' = (\<lambda> x. False)"  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
48  | 
code_pred (expected_modes: o => bool, i => bool) [inductify] EmptyPred' .  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) EmptyRel .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
54  | 
inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
code_pred  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
58  | 
(expected_modes: (o => o => bool) => o => o => bool, (o => o => bool) => i => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
59  | 
(o => o => bool) => o => i => bool, (o => o => bool) => i => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
60  | 
(i => o => bool) => o => o => bool, (i => o => bool) => i => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
61  | 
(i => o => bool) => o => i => bool, (i => o => bool) => i => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
62  | 
(o => i => bool) => o => o => bool, (o => i => bool) => i => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
63  | 
(o => i => bool) => o => i => bool, (o => i => bool) => i => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
64  | 
(i => i => bool) => o => o => bool, (i => i => bool) => i => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
65  | 
(i => i => bool) => o => i => bool, (i => i => bool) => i => i => bool)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
66  | 
EmptyClosure .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
68  | 
thm EmptyClosure.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
70  | 
(* TODO: inductive package is broken!  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
71  | 
inductive False'' :: "bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
72  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
73  | 
"False \<Longrightarrow> False''"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
74  | 
|
| 40100 | 75  | 
code_pred (expected_modes: bool) False'' .  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
77  | 
inductive EmptySet'' :: "'a \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
79  | 
"False \<Longrightarrow> EmptySet'' x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
80  | 
|
| 40100 | 81  | 
code_pred (expected_modes: i => bool, o => bool) [inductify] EmptySet'' .  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
82  | 
*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
84  | 
consts a' :: 'a  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
85  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
86  | 
inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
87  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
88  | 
"Fact a' a'"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
90  | 
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
91  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
92  | 
inductive zerozero :: "nat * nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
94  | 
"zerozero (0, 0)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
95  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
96  | 
code_pred (expected_modes: i => bool, i * o => bool, o * i => bool, o => bool) zerozero .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
97  | 
code_pred [dseq] zerozero .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
98  | 
code_pred [random_dseq] zerozero .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
100  | 
thm zerozero.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
101  | 
thm zerozero.dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
102  | 
thm zerozero.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
104  | 
text {* We expect the user to expand the tuples in the values command.
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
105  | 
The following values command is not supported. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
106  | 
(*values "{x. zerozero x}" *)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
107  | 
text {* Instead, the user must type *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
108  | 
values "{(x, y). zerozero (x, y)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
110  | 
values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
111  | 
values [expected "{(0::nat, 0::nat)}" dseq 1] "{(x, y). zerozero (x, y)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
112  | 
values [expected "{(0::nat, 0::nat)}" dseq 2] "{(x, y). zerozero (x, y)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
113  | 
values [expected "{}" random_dseq 1, 1, 2] "{(x, y). zerozero (x, y)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
114  | 
values [expected "{(0::nat, 0:: nat)}" random_dseq 1, 1, 3] "{(x, y). zerozero (x, y)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
115  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
116  | 
inductive nested_tuples :: "((int * int) * int * int) => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
117  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
118  | 
"nested_tuples ((0, 1), 2, 3)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
120  | 
code_pred nested_tuples .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
121  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
122  | 
inductive JamesBond :: "nat => int => natural => bool"  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
123  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
124  | 
"JamesBond 0 0 7"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
125  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
126  | 
code_pred JamesBond .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
127  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
128  | 
values [expected "{(0::nat, 0::int , 7::natural)}"] "{(a, b, c). JamesBond a b c}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
129  | 
values [expected "{(0::nat, 7::natural, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
130  | 
values [expected "{(0::int, 0::nat, 7::natural)}"] "{(b, a, c). JamesBond a b c}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
131  | 
values [expected "{(0::int, 7::natural, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
132  | 
values [expected "{(7::natural, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
133  | 
values [expected "{(7::natural, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
134  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
135  | 
values [expected "{(7::natural, 0::int)}"] "{(a, b). JamesBond 0 b a}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
136  | 
values [expected "{(7::natural, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
47433 
diff
changeset
 | 
137  | 
values [expected "{(0::nat, 7::natural)}"] "{(a, c). JamesBond a 0 c}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
140  | 
subsection {* Alternative Rules *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
141  | 
|
| 58310 | 142  | 
datatype char = C | D | E | F | G | H  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
143  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
144  | 
inductive is_C_or_D  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
145  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
146  | 
"(x = C) \<or> (x = D) ==> is_C_or_D x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
147  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
148  | 
code_pred (expected_modes: i => bool) is_C_or_D .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
149  | 
thm is_C_or_D.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
150  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
151  | 
inductive is_D_or_E  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
152  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
153  | 
"(x = D) \<or> (x = E) ==> is_D_or_E x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
155  | 
lemma [code_pred_intro]:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
156  | 
"is_D_or_E D"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
157  | 
by (auto intro: is_D_or_E.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
158  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
159  | 
lemma [code_pred_intro]:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
160  | 
"is_D_or_E E"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
161  | 
by (auto intro: is_D_or_E.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
162  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
163  | 
code_pred (expected_modes: o => bool, i => bool) is_D_or_E  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
164  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
165  | 
case is_D_or_E  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
166  | 
from is_D_or_E.prems show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
167  | 
proof  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
168  | 
fix xa  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
169  | 
assume x: "x = xa"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
170  | 
assume "xa = D \<or> xa = E"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
171  | 
from this show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
172  | 
proof  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
173  | 
assume "xa = D" from this x is_D_or_E(1) show thesis by simp  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
174  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
175  | 
assume "xa = E" from this x is_D_or_E(2) show thesis by simp  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
176  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
177  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
178  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
180  | 
thm is_D_or_E.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
181  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
182  | 
inductive is_F_or_G  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
183  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
184  | 
"x = F \<or> x = G ==> is_F_or_G x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
185  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
186  | 
lemma [code_pred_intro]:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
187  | 
"is_F_or_G F"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
188  | 
by (auto intro: is_F_or_G.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
189  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
190  | 
lemma [code_pred_intro]:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
191  | 
"is_F_or_G G"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
192  | 
by (auto intro: is_F_or_G.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
193  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
194  | 
inductive is_FGH  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
195  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
196  | 
"is_F_or_G x ==> is_FGH x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
197  | 
| "is_FGH H"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
199  | 
text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
200  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
201  | 
code_pred (expected_modes: o => bool, i => bool) is_FGH  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
202  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
203  | 
case is_F_or_G  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
204  | 
from is_F_or_G.prems show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
205  | 
proof  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
206  | 
fix xa  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
207  | 
assume x: "x = xa"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
208  | 
assume "xa = F \<or> xa = G"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
209  | 
from this show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
210  | 
proof  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
211  | 
assume "xa = F"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
212  | 
from this x is_F_or_G(1) show thesis by simp  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
213  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
214  | 
assume "xa = G"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
215  | 
from this x is_F_or_G(2) show thesis by simp  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
216  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
217  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
218  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
220  | 
subsection {* Named alternative rules *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
222  | 
inductive appending  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
223  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
224  | 
nil: "appending [] ys ys"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
225  | 
| cons: "appending xs ys zs \<Longrightarrow> appending (x#xs) ys (x#zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
226  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
227  | 
lemma appending_alt_nil: "ys = zs \<Longrightarrow> appending [] ys zs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
228  | 
by (auto intro: appending.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
229  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
230  | 
lemma appending_alt_cons: "xs' = x # xs \<Longrightarrow> appending xs ys zs \<Longrightarrow> zs' = x # zs \<Longrightarrow> appending xs' ys zs'"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
231  | 
by (auto intro: appending.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
232  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
233  | 
text {* With code_pred_intro, we can give fact names to the alternative rules,
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
234  | 
which are used for the code_pred command. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
235  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
236  | 
declare appending_alt_nil[code_pred_intro alt_nil] appending_alt_cons[code_pred_intro alt_cons]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
237  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
238  | 
code_pred appending  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
239  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
240  | 
case appending  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
241  | 
from appending.prems show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
242  | 
proof(cases)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
243  | 
case nil  | 
| 60565 | 244  | 
from appending.alt_nil nil show thesis by auto  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
245  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
246  | 
case cons  | 
| 60565 | 247  | 
from appending.alt_cons cons show thesis by fastforce  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
248  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
249  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
250  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
251  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
252  | 
inductive ya_even and ya_odd :: "nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
253  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
254  | 
even_zero: "ya_even 0"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
255  | 
| odd_plus1: "ya_even x ==> ya_odd (x + 1)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
256  | 
| even_plus1: "ya_odd x ==> ya_even (x + 1)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
257  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
258  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
259  | 
declare even_zero[code_pred_intro even_0]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
260  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
261  | 
lemma [code_pred_intro odd_Suc]: "ya_even x ==> ya_odd (Suc x)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
262  | 
by (auto simp only: Suc_eq_plus1 intro: ya_even_ya_odd.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
263  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
264  | 
lemma [code_pred_intro even_Suc]:"ya_odd x ==> ya_even (Suc x)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
265  | 
by (auto simp only: Suc_eq_plus1 intro: ya_even_ya_odd.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
266  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
267  | 
code_pred ya_even  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
268  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
269  | 
case ya_even  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
270  | 
from ya_even.prems show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
271  | 
proof (cases)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
272  | 
case even_zero  | 
| 60565 | 273  | 
from even_zero ya_even.even_0 show thesis by simp  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
274  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
275  | 
case even_plus1  | 
| 60565 | 276  | 
from even_plus1 ya_even.even_Suc show thesis by simp  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
277  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
278  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
279  | 
case ya_odd  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
280  | 
from ya_odd.prems show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
281  | 
proof (cases)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
282  | 
case odd_plus1  | 
| 60565 | 283  | 
from odd_plus1 ya_odd.odd_Suc show thesis by simp  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
284  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
285  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
286  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
287  | 
subsection {* Preprocessor Inlining  *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
288  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
289  | 
definition "equals == (op =)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
290  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
291  | 
inductive zerozero' :: "nat * nat => bool" where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
292  | 
"equals (x, y) (0, 0) ==> zerozero' (x, y)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
293  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
294  | 
code_pred (expected_modes: i => bool) zerozero' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
295  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
296  | 
lemma zerozero'_eq: "zerozero' x == zerozero x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
297  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
298  | 
have "zerozero' = zerozero"  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
299  | 
apply (auto simp add: fun_eq_iff)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
300  | 
apply (cases rule: zerozero'.cases)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
301  | 
apply (auto simp add: equals_def intro: zerozero.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
302  | 
apply (cases rule: zerozero.cases)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
303  | 
apply (auto simp add: equals_def intro: zerozero'.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
304  | 
done  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
305  | 
from this show "zerozero' x == zerozero x" by auto  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
306  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
307  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
308  | 
declare zerozero'_eq [code_pred_inline]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
309  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
310  | 
definition "zerozero'' x == zerozero' x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
312  | 
text {* if preprocessing fails, zerozero'' will not have all modes. *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
313  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
314  | 
code_pred (expected_modes: i * i => bool, i * o => bool, o * i => bool, o => bool) [inductify] zerozero'' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
315  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
316  | 
subsection {* Sets *}
 | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
317  | 
(*  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
318  | 
inductive_set EmptySet :: "'a set"  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
319  | 
|
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
320  | 
code_pred (expected_modes: o => bool, i => bool) EmptySet .  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
321  | 
|
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
322  | 
definition EmptySet' :: "'a set"  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
323  | 
where "EmptySet' = {}"
 | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
324  | 
|
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
325  | 
code_pred (expected_modes: o => bool, i => bool) [inductify] EmptySet' .  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
326  | 
*)  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
327  | 
subsection {* Numerals *}
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
328  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
329  | 
definition  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
330  | 
"one_or_two = (%x. x = Suc 0 \<or> ( x = Suc (Suc 0)))"  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
331  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
332  | 
code_pred [inductify] one_or_two .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
333  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
334  | 
code_pred [dseq] one_or_two .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
335  | 
code_pred [random_dseq] one_or_two .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
336  | 
thm one_or_two.dseq_equation  | 
| 51144 | 337  | 
values [expected "{Suc 0, Suc (Suc 0)}"] "{x. one_or_two x}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
338  | 
values [random_dseq 0,0,10] 3 "{x. one_or_two x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
339  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
340  | 
inductive one_or_two' :: "nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
341  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
342  | 
"one_or_two' 1"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
343  | 
| "one_or_two' 2"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
344  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
345  | 
code_pred one_or_two' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
346  | 
thm one_or_two'.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
347  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
348  | 
values "{x. one_or_two' x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
349  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
350  | 
definition one_or_two'':  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
351  | 
"one_or_two'' == (%x. x = 1 \<or> x = (2::nat))"  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
352  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
353  | 
code_pred [inductify] one_or_two'' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
354  | 
thm one_or_two''.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
355  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
356  | 
values "{x. one_or_two'' x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
357  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
358  | 
subsection {* even predicate *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
359  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
360  | 
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
361  | 
"even 0"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
362  | 
| "even n \<Longrightarrow> odd (Suc n)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
363  | 
| "odd n \<Longrightarrow> even (Suc n)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
364  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
365  | 
code_pred (expected_modes: i => bool, o => bool) even .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
366  | 
code_pred [dseq] even .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
367  | 
code_pred [random_dseq] even .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
368  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
369  | 
thm odd.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
370  | 
thm even.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
371  | 
thm odd.dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
372  | 
thm even.dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
373  | 
thm odd.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
374  | 
thm even.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
375  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
376  | 
values "{x. even 2}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
377  | 
values "{x. odd 2}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
378  | 
values 10 "{n. even n}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
379  | 
values 10 "{n. odd n}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
380  | 
values [expected "{}" dseq 2] "{x. even 6}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
381  | 
values [expected "{}" dseq 6] "{x. even 6}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
382  | 
values [expected "{()}" dseq 7] "{x. even 6}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
383  | 
values [dseq 2] "{x. odd 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
384  | 
values [dseq 6] "{x. odd 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
385  | 
values [dseq 7] "{x. odd 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
386  | 
values [expected "{()}" dseq 8] "{x. odd 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
387  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
388  | 
values [expected "{}" dseq 0] 8 "{x. even x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
389  | 
values [expected "{0::nat}" dseq 1] 8 "{x. even x}"
 | 
| 51144 | 390  | 
values [expected "{0, Suc (Suc 0)}" dseq 3] 8 "{x. even x}"
 | 
391  | 
values [expected "{0, Suc (Suc 0)}" dseq 4] 8 "{x. even x}"
 | 
|
392  | 
values [expected "{0, Suc (Suc 0), Suc (Suc (Suc (Suc 0)))}" dseq 6] 8 "{x. even x}"
 | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
393  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
394  | 
values [random_dseq 1, 1, 0] 8 "{x. even x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
395  | 
values [random_dseq 1, 1, 1] 8 "{x. even x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
396  | 
values [random_dseq 1, 1, 2] 8 "{x. even x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
397  | 
values [random_dseq 1, 1, 3] 8 "{x. even x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
398  | 
values [random_dseq 1, 1, 6] 8 "{x. even x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
399  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
400  | 
values [expected "{}" random_dseq 1, 1, 7] "{x. odd 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
401  | 
values [random_dseq 1, 1, 8] "{x. odd 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
402  | 
values [random_dseq 1, 1, 9] "{x. odd 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
403  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
404  | 
definition odd' where "odd' x == \<not> even x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
405  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
406  | 
code_pred (expected_modes: i => bool) [inductify] odd' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
407  | 
code_pred [dseq inductify] odd' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
408  | 
code_pred [random_dseq inductify] odd' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
409  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
410  | 
values [expected "{}" dseq 2] "{x. odd' 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
411  | 
values [expected "{()}" dseq 9] "{x. odd' 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
412  | 
values [expected "{}" dseq 2] "{x. odd' 8}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
413  | 
values [expected "{}" dseq 10] "{x. odd' 8}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
414  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
415  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
416  | 
inductive is_even :: "nat \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
417  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
418  | 
"n mod 2 = 0 \<Longrightarrow> is_even n"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
419  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
420  | 
code_pred (expected_modes: i => bool) is_even .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
421  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
422  | 
subsection {* append predicate *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
423  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
424  | 
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
425  | 
"append [] xs xs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
426  | 
| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
427  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
428  | 
code_pred (modes: i => i => o => bool as "concat", o => o => i => bool as "slice", o => i => i => bool as prefix,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
429  | 
i => o => i => bool as suffix, i => i => i => bool) append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
430  | 
code_pred (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as "concat", o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as "slice", o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool as prefix,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
431  | 
i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool) append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
432  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
433  | 
code_pred [dseq] append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
434  | 
code_pred [random_dseq] append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
435  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
436  | 
thm append.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
437  | 
thm append.dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
438  | 
thm append.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
439  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
440  | 
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
441  | 
values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
442  | 
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
443  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
444  | 
values [expected "{}" dseq 0] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 51144 | 445  | 
values [expected "{(([]::nat list), [Suc 0, Suc (Suc 0), Suc (Suc (Suc 0)), Suc (Suc (Suc (Suc 0))), Suc (Suc (Suc (Suc (Suc 0))))])}" dseq 1] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
446  | 
values [dseq 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
447  | 
values [dseq 6] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
448  | 
values [random_dseq 1, 1, 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
449  | 
values [random_dseq 1, 1, 1] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
450  | 
values [random_dseq 1, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
451  | 
values [random_dseq 3, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
452  | 
values [random_dseq 1, 3, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
453  | 
values [random_dseq 1, 1, 4] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
454  | 
|
| 56927 | 455  | 
value "Predicate.the (concat [0::int, 1, 2] [3, 4, 5])"  | 
456  | 
value "Predicate.the (slice ([]::int list))"  | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
457  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
458  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
459  | 
text {* tricky case with alternative rules *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
460  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
461  | 
inductive append2  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
462  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
463  | 
"append2 [] xs xs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
464  | 
| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
465  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
466  | 
lemma append2_Nil: "append2 [] (xs::'b list) xs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
467  | 
by (simp add: append2.intros(1))  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
468  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
469  | 
lemmas [code_pred_intro] = append2_Nil append2.intros(2)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
470  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
471  | 
code_pred (expected_modes: i => i => o => bool, o => o => i => bool, o => i => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
472  | 
i => o => i => bool, i => i => i => bool) append2  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
473  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
474  | 
case append2  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
475  | 
from append2.prems show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
476  | 
proof  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
477  | 
fix xs  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
478  | 
assume "xa = []" "xb = xs" "xc = xs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
479  | 
from this append2(1) show thesis by simp  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
480  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
481  | 
fix xs ys zs x  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
482  | 
assume "xa = x # xs" "xb = ys" "xc = x # zs" "append2 xs ys zs"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42463 
diff
changeset
 | 
483  | 
from this append2(2) show thesis by fastforce  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
484  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
485  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
486  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
487  | 
inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
488  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
489  | 
"tupled_append ([], xs, xs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
490  | 
| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
491  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
492  | 
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
493  | 
i * o * i => bool, i * i * i => bool) tupled_append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
494  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
495  | 
code_pred (expected_modes: i \<times> i \<times> o \<Rightarrow> bool, o \<times> o \<times> i \<Rightarrow> bool, o \<times> i \<times> i \<Rightarrow> bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
496  | 
i \<times> o \<times> i \<Rightarrow> bool, i \<times> i \<times> i \<Rightarrow> bool) tupled_append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
497  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
498  | 
code_pred [random_dseq] tupled_append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
499  | 
thm tupled_append.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
500  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
501  | 
values "{xs. tupled_append ([(1::nat), 2, 3], [4, 5], xs)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
502  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
503  | 
inductive tupled_append'  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
504  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
505  | 
"tupled_append' ([], xs, xs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
506  | 
| "[| ys = fst (xa, y); x # zs = snd (xa, y);  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
507  | 
tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
508  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
509  | 
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
510  | 
i * o * i => bool, i * i * i => bool) tupled_append' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
511  | 
thm tupled_append'.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
512  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
513  | 
inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
514  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
515  | 
"tupled_append'' ([], xs, xs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
516  | 
| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
517  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
518  | 
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
519  | 
i * o * i => bool, i * i * i => bool) tupled_append'' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
520  | 
thm tupled_append''.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
521  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
522  | 
inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
523  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
524  | 
"tupled_append''' ([], xs, xs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
525  | 
| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
526  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
527  | 
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
528  | 
i * o * i => bool, i * i * i => bool) tupled_append''' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
529  | 
thm tupled_append'''.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
530  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
531  | 
subsection {* map_ofP predicate *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
532  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
533  | 
inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
534  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
535  | 
"map_ofP ((a, b)#xs) a b"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
536  | 
| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
537  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
538  | 
code_pred (expected_modes: i => o => o => bool, i => i => o => bool, i => o => i => bool, i => i => i => bool) map_ofP .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
539  | 
thm map_ofP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
540  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
541  | 
subsection {* filter predicate *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
542  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
543  | 
inductive filter1  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
544  | 
for P  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
545  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
546  | 
"filter1 P [] []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
547  | 
| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
548  | 
| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
549  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
550  | 
code_pred (expected_modes: (i => bool) => i => o => bool, (i => bool) => i => i => bool) filter1 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
551  | 
code_pred [dseq] filter1 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
552  | 
code_pred [random_dseq] filter1 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
553  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
554  | 
thm filter1.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
555  | 
|
| 51144 | 556  | 
values [expected "{[0, Suc (Suc 0), Suc (Suc (Suc (Suc 0)))]}"] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
557  | 
values [expected "{}" dseq 9] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
 | 
| 51144 | 558  | 
values [expected "{[0, Suc (Suc 0), Suc (Suc (Suc (Suc 0)))]}" dseq 10] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
559  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
560  | 
inductive filter2  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
561  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
562  | 
"filter2 P [] []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
563  | 
| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
564  | 
| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
565  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
566  | 
code_pred (expected_modes: (i => bool) => i => i => bool, (i => bool) => i => o => bool) filter2 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
567  | 
code_pred [dseq] filter2 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
568  | 
code_pred [random_dseq] filter2 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
569  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
570  | 
thm filter2.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
571  | 
thm filter2.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
572  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
573  | 
inductive filter3  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
574  | 
for P  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
575  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
576  | 
"List.filter P xs = ys ==> filter3 P xs ys"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
577  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
578  | 
code_pred (expected_modes: (o => bool) => i => o => bool, (o => bool) => i => i => bool , (i => bool) => i => o => bool, (i => bool) => i => i => bool) [skip_proof] filter3 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
579  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
580  | 
code_pred filter3 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
581  | 
thm filter3.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
582  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
583  | 
(*  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
584  | 
inductive filter4  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
585  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
586  | 
"List.filter P xs = ys ==> filter4 P xs ys"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
587  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
588  | 
code_pred (expected_modes: i => i => o => bool, i => i => i => bool) filter4 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
589  | 
(*code_pred [depth_limited] filter4 .*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
590  | 
(*code_pred [random] filter4 .*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
591  | 
*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
592  | 
subsection {* reverse predicate *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
593  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
594  | 
inductive rev where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
595  | 
"rev [] []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
596  | 
| "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
597  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
598  | 
code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) rev .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
599  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
600  | 
thm rev.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
601  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
602  | 
values "{xs. rev [0, 1, 2, 3::nat] xs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
603  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
604  | 
inductive tupled_rev where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
605  | 
"tupled_rev ([], [])"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
606  | 
| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
607  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
608  | 
code_pred (expected_modes: i * o => bool, o * i => bool, i * i => bool) tupled_rev .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
609  | 
thm tupled_rev.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
610  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
611  | 
subsection {* partition predicate *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
612  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
613  | 
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
614  | 
for f where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
615  | 
"partition f [] [] []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
616  | 
| "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
617  | 
| "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
618  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
619  | 
code_pred (expected_modes: (i => bool) => i => o => o => bool, (i => bool) => o => i => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
620  | 
(i => bool) => i => i => o => bool, (i => bool) => i => o => i => bool, (i => bool) => i => i => i => bool)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
621  | 
partition .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
622  | 
code_pred [dseq] partition .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
623  | 
code_pred [random_dseq] partition .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
624  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
625  | 
values 10 "{(ys, zs). partition is_even
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
626  | 
[0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
627  | 
values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
628  | 
values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
629  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
630  | 
inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
631  | 
for f where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
632  | 
"tupled_partition f ([], [], [])"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
633  | 
| "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
634  | 
| "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
635  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
636  | 
code_pred (expected_modes: (i => bool) => i => bool, (i => bool) => (i * i * o) => bool, (i => bool) => (i * o * i) => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
637  | 
(i => bool) => (o * i * i) => bool, (i => bool) => (i * o * o) => bool) tupled_partition .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
638  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
639  | 
thm tupled_partition.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
640  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
641  | 
lemma [code_pred_intro]:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
642  | 
"r a b \<Longrightarrow> tranclp r a b"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
643  | 
"r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
644  | 
by auto  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
645  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
646  | 
subsection {* transitive predicate *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
647  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
648  | 
text {* Also look at the tabled transitive closure in the Library *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
649  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
650  | 
code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
651  | 
(o => i => bool) => i => i => bool, (o => i => bool) => o => i => bool as backwards_trancl, (o => o => bool) => i => i => bool, (o => o => bool) => i => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
652  | 
(o => o => bool) => o => i => bool, (o => o => bool) => o => o => bool) tranclp  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
653  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
654  | 
case tranclp  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
655  | 
from this converse_tranclpE[OF tranclp.prems] show thesis by metis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
656  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
657  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
658  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
659  | 
code_pred [dseq] tranclp .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
660  | 
code_pred [random_dseq] tranclp .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
661  | 
thm tranclp.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
662  | 
thm tranclp.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
663  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
664  | 
inductive rtrancl' :: "'a => 'a => ('a => 'a => bool) => bool" 
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
665  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
666  | 
"rtrancl' x x r"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
667  | 
| "r x y ==> rtrancl' y z r ==> rtrancl' x z r"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
668  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
669  | 
code_pred [random_dseq] rtrancl' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
670  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
671  | 
thm rtrancl'.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
672  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
673  | 
inductive rtrancl'' :: "('a * 'a * ('a \<Rightarrow> 'a \<Rightarrow> bool)) \<Rightarrow> bool"  
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
674  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
675  | 
"rtrancl'' (x, x, r)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
676  | 
| "r x y \<Longrightarrow> rtrancl'' (y, z, r) \<Longrightarrow> rtrancl'' (x, z, r)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
677  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
678  | 
code_pred rtrancl'' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
679  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
680  | 
inductive rtrancl''' :: "('a * ('a * 'a) * ('a * 'a => bool)) => bool" 
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
681  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
682  | 
"rtrancl''' (x, (x, x), r)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
683  | 
| "r (x, y) ==> rtrancl''' (y, (z, z), r) ==> rtrancl''' (x, (z, z), r)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
684  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
685  | 
code_pred rtrancl''' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
686  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
687  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
688  | 
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
689  | 
"succ 0 1"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
690  | 
| "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
691  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
692  | 
code_pred (modes: i => i => bool, i => o => bool, o => i => bool, o => o => bool) succ .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
693  | 
code_pred [random_dseq] succ .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
694  | 
thm succ.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
695  | 
thm succ.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
696  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
697  | 
values 10 "{(m, n). succ n m}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
698  | 
values "{m. succ 0 m}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
699  | 
values "{m. succ m 0}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
700  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
701  | 
text {* values command needs mode annotation of the parameter succ
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
702  | 
to disambiguate which mode is to be chosen. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
703  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
704  | 
values [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
705  | 
values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
706  | 
values 20 "{(n, m). tranclp succ n m}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
707  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
708  | 
inductive example_graph :: "int => int => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
709  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
710  | 
"example_graph 0 1"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
711  | 
| "example_graph 1 2"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
712  | 
| "example_graph 1 3"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
713  | 
| "example_graph 4 7"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
714  | 
| "example_graph 4 5"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
715  | 
| "example_graph 5 6"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
716  | 
| "example_graph 7 6"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
717  | 
| "example_graph 7 8"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
718  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
719  | 
inductive not_reachable_in_example_graph :: "int => int => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
720  | 
where "\<not> (tranclp example_graph x y) ==> not_reachable_in_example_graph x y"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
721  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
722  | 
code_pred (expected_modes: i => i => bool) not_reachable_in_example_graph .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
723  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
724  | 
thm not_reachable_in_example_graph.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
725  | 
thm tranclp.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
726  | 
value "not_reachable_in_example_graph 0 3"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
727  | 
value "not_reachable_in_example_graph 4 8"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
728  | 
value "not_reachable_in_example_graph 5 6"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
729  | 
text {* rtrancl compilation is strange! *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
730  | 
(*  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
731  | 
value "not_reachable_in_example_graph 0 4"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
732  | 
value "not_reachable_in_example_graph 1 6"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
733  | 
value "not_reachable_in_example_graph 8 4"*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
734  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
735  | 
code_pred [dseq] not_reachable_in_example_graph .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
736  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
737  | 
values [dseq 6] "{x. tranclp example_graph 0 3}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
738  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
739  | 
values [dseq 0] "{x. not_reachable_in_example_graph 0 3}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
740  | 
values [dseq 0] "{x. not_reachable_in_example_graph 0 4}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
741  | 
values [dseq 20] "{x. not_reachable_in_example_graph 0 4}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
742  | 
values [dseq 6] "{x. not_reachable_in_example_graph 0 3}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
743  | 
values [dseq 3] "{x. not_reachable_in_example_graph 4 2}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
744  | 
values [dseq 6] "{x. not_reachable_in_example_graph 4 2}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
745  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
746  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
747  | 
inductive not_reachable_in_example_graph' :: "int => int => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
748  | 
where "\<not> (rtranclp example_graph x y) ==> not_reachable_in_example_graph' x y"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
749  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
750  | 
code_pred not_reachable_in_example_graph' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
751  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
752  | 
value "not_reachable_in_example_graph' 0 3"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
753  | 
(* value "not_reachable_in_example_graph' 0 5" would not terminate *)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
754  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
755  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
756  | 
(*values [depth_limited 0] "{x. not_reachable_in_example_graph' 0 3}"*)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
757  | 
(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
758  | 
(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"*)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
759  | 
(*values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
760  | 
(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
761  | 
(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
762  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
763  | 
code_pred [dseq] not_reachable_in_example_graph' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
764  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
765  | 
(*thm not_reachable_in_example_graph'.dseq_equation*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
766  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
767  | 
(*values [dseq 0] "{x. not_reachable_in_example_graph' 0 3}"*)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
768  | 
(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
769  | 
(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
770  | 
values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
771  | 
(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
772  | 
(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
773  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
774  | 
subsection {* Free function variable *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
775  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
776  | 
inductive FF :: "nat => nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
777  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
778  | 
"f x = y ==> FF x y"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
779  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
780  | 
code_pred FF .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
781  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
782  | 
subsection {* IMP *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
783  | 
|
| 42463 | 784  | 
type_synonym var = nat  | 
785  | 
type_synonym state = "int list"  | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
786  | 
|
| 58310 | 787  | 
datatype com =  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
788  | 
Skip |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
789  | 
Ass var "state => int" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
790  | 
Seq com com |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
791  | 
IF "state => bool" com com |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
792  | 
While "state => bool" com  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
793  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
794  | 
inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
795  | 
"tupled_exec (Skip, s, s)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
796  | 
"tupled_exec (Ass x e, s, s[x := e(s)])" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
797  | 
"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
798  | 
"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
799  | 
"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
800  | 
"~b s ==> tupled_exec (While b c, s, s)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
801  | 
"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
802  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
803  | 
code_pred tupled_exec .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
804  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
805  | 
values "{s. tupled_exec (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))), [3, 5], s)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
806  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
807  | 
subsection {* CCS *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
808  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
809  | 
text{* This example formalizes finite CCS processes without communication or
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
810  | 
recursion. For simplicity, labels are natural numbers. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
811  | 
|
| 58310 | 812  | 
datatype proc = nil | pre nat proc | or proc proc | par proc proc  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
813  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
814  | 
inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
815  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
816  | 
"tupled_step (pre n p, n, p)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
817  | 
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
818  | 
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
819  | 
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
820  | 
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
821  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
822  | 
code_pred tupled_step .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
823  | 
thm tupled_step.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
824  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
825  | 
subsection {* divmod *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
826  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
827  | 
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
828  | 
"k < l \<Longrightarrow> divmod_rel k l 0 k"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
829  | 
| "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
830  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
831  | 
code_pred divmod_rel .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
832  | 
thm divmod_rel.equation  | 
| 56927 | 833  | 
value "Predicate.the (divmod_rel_i_i_o_o 1705 42)"  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
834  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
835  | 
subsection {* Transforming predicate logic into logic programs *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
836  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
837  | 
subsection {* Transforming functions into logic programs *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
838  | 
definition  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
839  | 
"case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
840  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
841  | 
code_pred [inductify, skip_proof] case_f .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
842  | 
thm case_fP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
843  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
844  | 
fun fold_map_idx where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
845  | 
"fold_map_idx f i y [] = (y, [])"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
846  | 
| "fold_map_idx f i y (x # xs) =  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
847  | 
(let (y', x') = f i y x; (y'', xs') = fold_map_idx f (Suc i) y' xs  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
848  | 
in (y'', x' # xs'))"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
849  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
850  | 
code_pred [inductify] fold_map_idx .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
851  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
852  | 
subsection {* Minimum *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
853  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
854  | 
definition Min  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
855  | 
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
856  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
857  | 
code_pred [inductify] Min .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
858  | 
thm Min.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
859  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
860  | 
subsection {* Lexicographic order *}
 | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
861  | 
text {* This example requires to handle the differences of sets and predicates in the predicate compiler,
 | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
862  | 
or to have a copy of all definitions on predicates due to the set-predicate distinction. *}  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
863  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
864  | 
(*  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
865  | 
declare lexord_def[code_pred_def]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
866  | 
code_pred [inductify] lexord .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
867  | 
code_pred [random_dseq inductify] lexord .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
868  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
869  | 
thm lexord.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
870  | 
thm lexord.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
871  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
872  | 
inductive less_than_nat :: "nat * nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
873  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
874  | 
"less_than_nat (0, x)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
875  | 
| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
876  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
877  | 
code_pred less_than_nat .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
878  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
879  | 
code_pred [dseq] less_than_nat .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
880  | 
code_pred [random_dseq] less_than_nat .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
881  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
882  | 
inductive test_lexord :: "nat list * nat list => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
883  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
884  | 
"lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
885  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
886  | 
code_pred test_lexord .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
887  | 
code_pred [dseq] test_lexord .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
888  | 
code_pred [random_dseq] test_lexord .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
889  | 
thm test_lexord.dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
890  | 
thm test_lexord.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
891  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
892  | 
values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
893  | 
(*values [depth_limited 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"*)
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
894  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
895  | 
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
896  | 
(*  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
897  | 
code_pred [inductify] lexn .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
898  | 
thm lexn.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
899  | 
*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
900  | 
(*  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
901  | 
code_pred [random_dseq inductify] lexn .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
902  | 
thm lexn.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
903  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
904  | 
values [random_dseq 4, 4, 6] 100 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
905  | 
*)  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
906  | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
907  | 
inductive has_length  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
908  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
909  | 
"has_length [] 0"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
910  | 
| "has_length xs i ==> has_length (x # xs) (Suc i)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
911  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
912  | 
lemma has_length:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
913  | 
"has_length xs n = (length xs = n)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
914  | 
proof (rule iffI)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
915  | 
assume "has_length xs n"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
916  | 
from this show "length xs = n"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
917  | 
by (rule has_length.induct) auto  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
918  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
919  | 
assume "length xs = n"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
920  | 
from this show "has_length xs n"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
921  | 
by (induct xs arbitrary: n) (auto intro: has_length.intros)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
922  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
923  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
924  | 
lemma lexn_intros [code_pred_intro]:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
925  | 
"has_length xs i ==> has_length ys i ==> r (x, y) ==> lexn r (Suc i) (x # xs, y # ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
926  | 
"lexn r i (xs, ys) ==> lexn r (Suc i) (x # xs, x # ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
927  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
928  | 
assume "has_length xs i" "has_length ys i" "r (x, y)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
929  | 
from this has_length show "lexn r (Suc i) (x # xs, y # ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
930  | 
unfolding lexn_conv Collect_def mem_def  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42463 
diff
changeset
 | 
931  | 
by fastforce  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
932  | 
next  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
933  | 
assume "lexn r i (xs, ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
934  | 
thm lexn_conv  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
935  | 
from this show "lexn r (Suc i) (x#xs, x#ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
936  | 
unfolding Collect_def mem_def lexn_conv  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
937  | 
apply auto  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
938  | 
apply (rule_tac x="x # xys" in exI)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
939  | 
by auto  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
940  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
941  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
942  | 
code_pred [random_dseq] lexn  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
943  | 
proof -  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
944  | 
fix r n xs ys  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
945  | 
assume 1: "lexn r n (xs, ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
946  | 
assume 2: "\<And>r' i x xs' y ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', y # ys') ==> has_length xs' i ==> has_length ys' i ==> r' (x, y) ==> thesis"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
947  | 
assume 3: "\<And>r' i x xs' ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', x # ys') ==> lexn r' i (xs', ys') ==> thesis"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
948  | 
from 1 2 3 show thesis  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
949  | 
unfolding lexn_conv Collect_def mem_def  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
950  | 
apply (auto simp add: has_length)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
951  | 
apply (case_tac xys)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
952  | 
apply auto  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42463 
diff
changeset
 | 
953  | 
apply fastforce  | 
| 
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42463 
diff
changeset
 | 
954  | 
apply fastforce done  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
955  | 
qed  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
956  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
957  | 
values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
958  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
959  | 
code_pred [inductify, skip_proof] lex .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
960  | 
thm lex.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
961  | 
thm lex_def  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
962  | 
declare lenlex_conv[code_pred_def]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
963  | 
code_pred [inductify, skip_proof] lenlex .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
964  | 
thm lenlex.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
965  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
966  | 
code_pred [random_dseq inductify] lenlex .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
967  | 
thm lenlex.random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
968  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
969  | 
values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
970  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
971  | 
thm lists.intros  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
972  | 
code_pred [inductify] lists .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
973  | 
thm lists.equation  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
974  | 
*)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
975  | 
subsection {* AVL Tree *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
976  | 
|
| 58310 | 977  | 
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
978  | 
fun height :: "'a tree => nat" where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
979  | 
"height ET = 0"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
980  | 
| "height (MKT x l r h) = max (height l) (height r) + 1"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
981  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
982  | 
primrec avl :: "'a tree => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
983  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
984  | 
"avl ET = True"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
985  | 
| "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and>  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
986  | 
h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
987  | 
(*  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
988  | 
code_pred [inductify] avl .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
989  | 
thm avl.equation*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
990  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
991  | 
code_pred [new_random_dseq inductify] avl .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
992  | 
thm avl.new_random_dseq_equation  | 
| 
40137
 
9eabcb1bfe50
changing test parameters in examples to get to a result within the global timelimit
 
bulwahn 
parents: 
40100 
diff
changeset
 | 
993  | 
(* TODO: has highly non-deterministic execution time!  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
994  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
995  | 
values [new_random_dseq 2, 1, 7] 5 "{t:: int tree. avl t}"
 | 
| 
40137
 
9eabcb1bfe50
changing test parameters in examples to get to a result within the global timelimit
 
bulwahn 
parents: 
40100 
diff
changeset
 | 
996  | 
*)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
997  | 
fun set_of  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
998  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
999  | 
"set_of ET = {}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1000  | 
| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1001  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1002  | 
fun is_ord :: "nat tree => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1003  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1004  | 
"is_ord ET = True"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1005  | 
| "is_ord (MKT n l r h) =  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1006  | 
((\<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)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1007  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1008  | 
(*  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1009  | 
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] set_of .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1010  | 
thm set_of.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1011  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1012  | 
code_pred (expected_modes: i => bool) [inductify] is_ord .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1013  | 
thm is_ord_aux.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1014  | 
thm is_ord.equation  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1015  | 
*)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1016  | 
subsection {* Definitions about Relations *}
 | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1017  | 
(*  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1018  | 
code_pred (modes:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1019  | 
(i * i => bool) => i * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1020  | 
(i * o => bool) => o * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1021  | 
(i * o => bool) => i * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1022  | 
(o * i => bool) => i * o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1023  | 
(o * i => bool) => i * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1024  | 
(o * o => bool) => o * o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1025  | 
(o * o => bool) => i * o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1026  | 
(o * o => bool) => o * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1027  | 
(o * o => bool) => i * i => bool) [inductify] converse .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1028  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1029  | 
thm converse.equation  | 
| 
47433
 
07f4bf913230
renamed "rel_comp" to "relcomp" (to be consistent with, e.g., "relpow")
 
griff 
parents: 
47108 
diff
changeset
 | 
1030  | 
code_pred [inductify] relcomp .  | 
| 
 
07f4bf913230
renamed "rel_comp" to "relcomp" (to be consistent with, e.g., "relpow")
 
griff 
parents: 
47108 
diff
changeset
 | 
1031  | 
thm relcomp.equation  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1032  | 
code_pred [inductify] Image .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1033  | 
thm Image.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1034  | 
declare singleton_iff[code_pred_inline]  | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44890 
diff
changeset
 | 
1035  | 
declare Id_on_def[unfolded Bex_def UNION_eq singleton_iff, code_pred_def]  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1036  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1037  | 
code_pred (expected_modes:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1038  | 
(o => bool) => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1039  | 
(o => bool) => i * o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1040  | 
(o => bool) => o * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1041  | 
(o => bool) => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1042  | 
(i => bool) => i * o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1043  | 
(i => bool) => o * i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1044  | 
(i => bool) => i => bool) [inductify] Id_on .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1045  | 
thm Id_on.equation  | 
| 
46752
 
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
 
haftmann 
parents: 
45970 
diff
changeset
 | 
1046  | 
thm Domain_unfold  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1047  | 
code_pred (modes:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1048  | 
(o * o => bool) => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1049  | 
(o * o => bool) => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1050  | 
(i * o => bool) => i => bool) [inductify] Domain .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1051  | 
thm Domain.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1052  | 
|
| 
46752
 
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
 
haftmann 
parents: 
45970 
diff
changeset
 | 
1053  | 
thm Domain_converse [symmetric]  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1054  | 
code_pred (modes:  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1055  | 
(o * o => bool) => o => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1056  | 
(o * o => bool) => i => bool,  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1057  | 
(o * i => bool) => i => bool) [inductify] Range .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1058  | 
thm Range.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1059  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1060  | 
code_pred [inductify] Field .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1061  | 
thm Field.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1062  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1063  | 
thm refl_on_def  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1064  | 
code_pred [inductify] refl_on .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1065  | 
thm refl_on.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1066  | 
code_pred [inductify] total_on .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1067  | 
thm total_on.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1068  | 
code_pred [inductify] antisym .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1069  | 
thm antisym.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1070  | 
code_pred [inductify] trans .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1071  | 
thm trans.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1072  | 
code_pred [inductify] single_valued .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1073  | 
thm single_valued.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1074  | 
thm inv_image_def  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1075  | 
code_pred [inductify] inv_image .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1076  | 
thm inv_image.equation  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1077  | 
*)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1078  | 
subsection {* Inverting list functions *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1079  | 
|
| 56679 | 1080  | 
code_pred [inductify, skip_proof] size_list' .  | 
1081  | 
code_pred [new_random_dseq inductify] size_list' .  | 
|
1082  | 
thm size_list'P.equation  | 
|
1083  | 
thm size_list'P.new_random_dseq_equation  | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1084  | 
|
| 56679 | 1085  | 
values [new_random_dseq 2,3,10] 3 "{xs. size_list'P (xs::nat list) (5::nat)}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1086  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1087  | 
code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) [inductify, skip_proof] List.concat .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1088  | 
thm concatP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1089  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1090  | 
values "{ys. concatP [[1, 2], [3, (4::int)]] ys}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1091  | 
values "{ys. concatP [[1, 2], [3]] [1, 2, (3::nat)]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1092  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1093  | 
code_pred [dseq inductify] List.concat .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1094  | 
thm concatP.dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1095  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1096  | 
values [dseq 3] 3  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1097  | 
  "{xs. concatP xs ([0] :: nat list)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1098  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1099  | 
values [dseq 5] 3  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1100  | 
  "{xs. concatP xs ([1] :: int list)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1101  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1102  | 
values [dseq 5] 3  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1103  | 
  "{xs. concatP xs ([1] :: nat list)}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1104  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1105  | 
values [dseq 5] 3  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1106  | 
  "{xs. concatP xs [(1::int), 2]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1107  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1108  | 
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] hd .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1109  | 
thm hdP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1110  | 
values "{x. hdP [1, 2, (3::int)] x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1111  | 
values "{(xs, x). hdP [1, 2, (3::int)] 1}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1112  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1113  | 
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] tl .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1114  | 
thm tlP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1115  | 
values "{x. tlP [1, 2, (3::nat)] x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1116  | 
values "{x. tlP [1, 2, (3::int)] [3]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1117  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1118  | 
code_pred [inductify, skip_proof] last .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1119  | 
thm lastP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1120  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1121  | 
code_pred [inductify, skip_proof] butlast .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1122  | 
thm butlastP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1123  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1124  | 
code_pred [inductify, skip_proof] take .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1125  | 
thm takeP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1126  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1127  | 
code_pred [inductify, skip_proof] drop .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1128  | 
thm dropP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1129  | 
code_pred [inductify, skip_proof] zip .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1130  | 
thm zipP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1131  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1132  | 
code_pred [inductify, skip_proof] upt .  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1133  | 
(*  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1134  | 
code_pred [inductify, skip_proof] remdups .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1135  | 
thm remdupsP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1136  | 
code_pred [dseq inductify] remdups .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1137  | 
values [dseq 4] 5 "{xs. remdupsP xs [1, (2::int)]}"
 | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1138  | 
*)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1139  | 
code_pred [inductify, skip_proof] remove1 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1140  | 
thm remove1P.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1141  | 
values "{xs. remove1P 1 xs [2, (3::int)]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1142  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1143  | 
code_pred [inductify, skip_proof] removeAll .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1144  | 
thm removeAllP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1145  | 
code_pred [dseq inductify] removeAll .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1146  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1147  | 
values [dseq 4] 10 "{xs. removeAllP 1 xs [(2::nat)]}"
 | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1148  | 
(*  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1149  | 
code_pred [inductify] distinct .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1150  | 
thm distinct.equation  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44928 
diff
changeset
 | 
1151  | 
*)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1152  | 
code_pred [inductify, skip_proof] replicate .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1153  | 
thm replicateP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1154  | 
values 5 "{(n, xs). replicateP n (0::int) xs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1155  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1156  | 
code_pred [inductify, skip_proof] splice .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1157  | 
thm splice.simps  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1158  | 
thm spliceP.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1159  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1160  | 
values "{xs. spliceP xs [1, 2, 3] [1, 1, 1, 2, 1, (3::nat)]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1161  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1162  | 
code_pred [inductify, skip_proof] List.rev .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1163  | 
code_pred [inductify] map .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1164  | 
code_pred [inductify] foldr .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1165  | 
code_pred [inductify] foldl .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1166  | 
code_pred [inductify] filter .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1167  | 
code_pred [random_dseq inductify] filter .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1168  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1169  | 
section {* Function predicate replacement *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1170  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1171  | 
text {*
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1172  | 
If the mode analysis uses the functional mode, we  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1173  | 
replace predicates that resulted from functions again by their functions.  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1174  | 
*}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1175  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1176  | 
inductive test_append  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1177  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1178  | 
"List.append xs ys = zs ==> test_append xs ys zs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1179  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1180  | 
code_pred [inductify, skip_proof] test_append .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1181  | 
thm test_append.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1182  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1183  | 
text {* If append is not turned into a predicate, then the mode
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1184  | 
o => o => i => bool could not be inferred. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1185  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1186  | 
values 4 "{(xs::int list, ys). test_append xs ys [1, 2, 3, 4]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1187  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1188  | 
text {* If appendP is not reverted back to a function, then mode i => i => o => bool
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1189  | 
fails after deleting the predicate equation. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1190  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1191  | 
declare appendP.equation[code del]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1192  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1193  | 
values "{xs::int list. test_append [1,2] [3,4] xs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1194  | 
values "{xs::int list. test_append (replicate 1000 1) (replicate 1000 2) xs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1195  | 
values "{xs::int list. test_append (replicate 2000 1) (replicate 2000 2) xs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1196  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1197  | 
text {* Redeclaring append.equation as code equation *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1198  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1199  | 
declare appendP.equation[code]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1200  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1201  | 
subsection {* Function with tuples *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1202  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1203  | 
fun append'  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1204  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1205  | 
"append' ([], ys) = ys"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1206  | 
| "append' (x # xs, ys) = x # append' (xs, ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1207  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1208  | 
inductive test_append'  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1209  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1210  | 
"append' (xs, ys) = zs ==> test_append' xs ys zs"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1211  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1212  | 
code_pred [inductify, skip_proof] test_append' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1213  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1214  | 
thm test_append'.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1215  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1216  | 
values "{(xs::int list, ys). test_append' xs ys [1, 2, 3, 4]}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1217  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1218  | 
declare append'P.equation[code del]  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1219  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1220  | 
values "{zs :: int list. test_append' [1,2,3] [4,5] zs}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1221  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1222  | 
section {* Arithmetic examples *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1223  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1224  | 
subsection {* Examples on nat *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1225  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1226  | 
inductive plus_nat_test :: "nat => nat => nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1227  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1228  | 
"x + y = z ==> plus_nat_test x y z"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1229  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1230  | 
code_pred [inductify, skip_proof] plus_nat_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1231  | 
code_pred [new_random_dseq inductify] plus_nat_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1232  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1233  | 
thm plus_nat_test.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1234  | 
thm plus_nat_test.new_random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1235  | 
|
| 51144 | 1236  | 
values [expected "{Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))}"] "{z. plus_nat_test 4 5 z}"
 | 
1237  | 
values [expected "{Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))}"] "{z. plus_nat_test 7 2 z}"
 | 
|
1238  | 
values [expected "{Suc (Suc (Suc (Suc 0)))}"] "{y. plus_nat_test 5 y 9}"
 | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1239  | 
values [expected "{}"] "{y. plus_nat_test 9 y 8}"
 | 
| 51144 | 1240  | 
values [expected "{Suc (Suc (Suc (Suc (Suc (Suc 0)))))}"] "{y. plus_nat_test 1 y 7}"
 | 
1241  | 
values [expected "{Suc (Suc 0)}"] "{x. plus_nat_test x 7 9}"
 | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1242  | 
values [expected "{}"] "{x. plus_nat_test x 9 7}"
 | 
| 51144 | 1243  | 
values [expected "{(0::nat, 0::nat)}"] "{(x, y). plus_nat_test x y 0}"
 | 
1244  | 
values [expected "{(0, Suc 0), (Suc 0, 0)}"] "{(x, y). plus_nat_test x y 1}"
 | 
|
1245  | 
values [expected "{(0, Suc (Suc (Suc (Suc (Suc 0))))),
 | 
|
1246  | 
(Suc 0, Suc (Suc (Suc (Suc 0)))),  | 
|
1247  | 
(Suc (Suc 0), Suc (Suc (Suc 0))),  | 
|
1248  | 
(Suc (Suc (Suc 0)), Suc (Suc 0)),  | 
|
1249  | 
(Suc (Suc (Suc (Suc 0))), Suc 0),  | 
|
1250  | 
(Suc (Suc (Suc (Suc (Suc 0)))), 0)}"]  | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1251  | 
  "{(x, y). plus_nat_test x y 5}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1252  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1253  | 
inductive minus_nat_test :: "nat => nat => nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1254  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1255  | 
"x - y = z ==> minus_nat_test x y z"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1256  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1257  | 
code_pred [inductify, skip_proof] minus_nat_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1258  | 
code_pred [new_random_dseq inductify] minus_nat_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1259  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1260  | 
thm minus_nat_test.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1261  | 
thm minus_nat_test.new_random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1262  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1263  | 
values [expected "{0::nat}"] "{z. minus_nat_test 4 5 z}"
 | 
| 51144 | 1264  | 
values [expected "{Suc (Suc (Suc (Suc (Suc 0))))}"] "{z. minus_nat_test 7 2 z}"
 | 
1265  | 
values [expected "{Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))))))))))}"] "{x. minus_nat_test x 7 9}"
 | 
|
1266  | 
values [expected "{Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))))))))))}"] "{x. minus_nat_test x 9 7}"
 | 
|
1267  | 
values [expected "{0, Suc 0, Suc (Suc 0), Suc (Suc (Suc 0))}"] "{x. minus_nat_test x 3 0}"
 | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1268  | 
values [expected "{0::nat}"] "{x. minus_nat_test x 0 0}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1269  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1270  | 
subsection {* Examples on int *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1271  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1272  | 
inductive plus_int_test :: "int => int => int => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1273  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1274  | 
"a + b = c ==> plus_int_test a b c"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1275  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1276  | 
code_pred [inductify, skip_proof] plus_int_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1277  | 
code_pred [new_random_dseq inductify] plus_int_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1278  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1279  | 
thm plus_int_test.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1280  | 
thm plus_int_test.new_random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1281  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1282  | 
values [expected "{1::int}"] "{a. plus_int_test a 6 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1283  | 
values [expected "{1::int}"] "{b. plus_int_test 6 b 7}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1284  | 
values [expected "{11::int}"] "{c. plus_int_test 5 6 c}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1285  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1286  | 
inductive minus_int_test :: "int => int => int => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1287  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1288  | 
"a - b = c ==> minus_int_test a b c"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1289  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1290  | 
code_pred [inductify, skip_proof] minus_int_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1291  | 
code_pred [new_random_dseq inductify] minus_int_test .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1292  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1293  | 
thm minus_int_test.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1294  | 
thm minus_int_test.new_random_dseq_equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1295  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1296  | 
values [expected "{4::int}"] "{c. minus_int_test 9 5 c}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1297  | 
values [expected "{9::int}"] "{a. minus_int_test a 4 5}"
 | 
| 
40885
 
da4bdafeef7c
adapted expected value to more idiomatic numeral representation
 
haftmann 
parents: 
40137 
diff
changeset
 | 
1298  | 
values [expected "{-1::int}"] "{b. minus_int_test 4 b 5}"
 | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1299  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1300  | 
subsection {* minus on bool *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1301  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1302  | 
inductive All :: "nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1303  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1304  | 
"All x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1305  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1306  | 
inductive None :: "nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1307  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1308  | 
definition "test_minus_bool x = (None x - All x)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1309  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1310  | 
code_pred [inductify] test_minus_bool .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1311  | 
thm test_minus_bool.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1312  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1313  | 
values "{x. test_minus_bool x}"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1314  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1315  | 
subsection {* Functions *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1316  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1317  | 
fun partial_hd :: "'a list => 'a option"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1318  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1319  | 
"partial_hd [] = Option.None"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1320  | 
| "partial_hd (x # xs) = Some x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1321  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1322  | 
inductive hd_predicate  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1323  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1324  | 
"partial_hd xs = Some x ==> hd_predicate xs x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1325  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1326  | 
code_pred (expected_modes: i => i => bool, i => o => bool) hd_predicate .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1327  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1328  | 
thm hd_predicate.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1329  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1330  | 
subsection {* Locales *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1331  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1332  | 
inductive hd_predicate2 :: "('a list => 'a option) => 'a list => 'a => bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1333  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1334  | 
"partial_hd' xs = Some x ==> hd_predicate2 partial_hd' xs x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1335  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1336  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1337  | 
thm hd_predicate2.intros  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1338  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1339  | 
code_pred (expected_modes: i => i => i => bool, i => i => o => bool) hd_predicate2 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1340  | 
thm hd_predicate2.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1341  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1342  | 
locale A = fixes partial_hd :: "'a list => 'a option" begin  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1343  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1344  | 
inductive hd_predicate_in_locale :: "'a list => 'a => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1345  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1346  | 
"partial_hd xs = Some x ==> hd_predicate_in_locale xs x"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1347  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1348  | 
end  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1349  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1350  | 
text {* The global introduction rules must be redeclared as introduction rules and then 
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1351  | 
one could invoke code_pred. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1352  | 
|
| 
39657
 
5e57675b7e40
moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
 
bulwahn 
parents: 
39655 
diff
changeset
 | 
1353  | 
declare A.hd_predicate_in_locale.intros [code_pred_intro]  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1354  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1355  | 
code_pred (expected_modes: i => i => i => bool, i => i => o => bool) A.hd_predicate_in_locale  | 
| 
39657
 
5e57675b7e40
moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
 
bulwahn 
parents: 
39655 
diff
changeset
 | 
1356  | 
by (auto elim: A.hd_predicate_in_locale.cases)  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1357  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1358  | 
interpretation A partial_hd .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1359  | 
thm hd_predicate_in_locale.intros  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1360  | 
text {* A locally compliant solution with a trivial interpretation fails, because
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1361  | 
the predicate compiler has very strict assumptions about the terms and their structure. *}  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1362  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1363  | 
(*code_pred hd_predicate_in_locale .*)  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1364  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1365  | 
section {* Integer example *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1366  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1367  | 
definition three :: int  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1368  | 
where "three = 3"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1369  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1370  | 
inductive is_three  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1371  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1372  | 
"is_three three"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1373  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1374  | 
code_pred is_three .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1375  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1376  | 
thm is_three.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1377  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1378  | 
section {* String.literal example *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1379  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1380  | 
definition Error_1  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1381  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1382  | 
"Error_1 = STR ''Error 1''"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1383  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1384  | 
definition Error_2  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1385  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1386  | 
"Error_2 = STR ''Error 2''"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1387  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1388  | 
inductive "is_error" :: "String.literal \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1389  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1390  | 
"is_error Error_1"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1391  | 
| "is_error Error_2"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1392  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1393  | 
code_pred is_error .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1394  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1395  | 
thm is_error.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1396  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1397  | 
inductive is_error' :: "String.literal \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1398  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1399  | 
"is_error' (STR ''Error1'')"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1400  | 
| "is_error' (STR ''Error2'')"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1401  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1402  | 
code_pred is_error' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1403  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1404  | 
thm is_error'.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1405  | 
|
| 58310 | 1406  | 
datatype ErrorObject = Error String.literal int  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1407  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1408  | 
inductive is_error'' :: "ErrorObject \<Rightarrow> bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1409  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1410  | 
"is_error'' (Error Error_1 3)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1411  | 
| "is_error'' (Error Error_2 4)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1412  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1413  | 
code_pred is_error'' .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1414  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1415  | 
thm is_error''.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1416  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1417  | 
section {* Another function example *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1418  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1419  | 
consts f :: "'a \<Rightarrow> 'a"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1420  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1421  | 
inductive fun_upd :: "(('a * 'b) * ('a \<Rightarrow> 'b)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1422  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1423  | 
"fun_upd ((x, a), s) (s(x := f a))"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1424  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1425  | 
code_pred fun_upd .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1426  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1427  | 
thm fun_upd.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1428  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1429  | 
section {* Examples for detecting switches *}
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1430  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1431  | 
inductive detect_switches1 where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1432  | 
"detect_switches1 [] []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1433  | 
| "detect_switches1 (x # xs) (y # ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1434  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1435  | 
code_pred [detect_switches, skip_proof] detect_switches1 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1436  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1437  | 
thm detect_switches1.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1438  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1439  | 
inductive detect_switches2 :: "('a => bool) => bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1440  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1441  | 
"detect_switches2 P"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1442  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1443  | 
code_pred [detect_switches, skip_proof] detect_switches2 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1444  | 
thm detect_switches2.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1445  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1446  | 
inductive detect_switches3 :: "('a => bool) => 'a list => bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1447  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1448  | 
"detect_switches3 P []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1449  | 
| "detect_switches3 P (x # xs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1450  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1451  | 
code_pred [detect_switches, skip_proof] detect_switches3 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1452  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1453  | 
thm detect_switches3.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1454  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1455  | 
inductive detect_switches4 :: "('a => bool) => 'a list => 'a list => bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1456  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1457  | 
"detect_switches4 P [] []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1458  | 
| "detect_switches4 P (x # xs) (y # ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1459  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1460  | 
code_pred [detect_switches, skip_proof] detect_switches4 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1461  | 
thm detect_switches4.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1462  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1463  | 
inductive detect_switches5 :: "('a => 'a => bool) => 'a list => 'a list => bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1464  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1465  | 
"detect_switches5 P [] []"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1466  | 
| "detect_switches5 P xs ys ==> P x y ==> detect_switches5 P (x # xs) (y # ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1467  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1468  | 
code_pred [detect_switches, skip_proof] detect_switches5 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1469  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1470  | 
thm detect_switches5.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1471  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1472  | 
inductive detect_switches6 :: "(('a => 'b => bool) * 'a list * 'b list) => bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1473  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1474  | 
"detect_switches6 (P, [], [])"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1475  | 
| "detect_switches6 (P, xs, ys) ==> P x y ==> detect_switches6 (P, x # xs, y # ys)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1476  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1477  | 
code_pred [detect_switches, skip_proof] detect_switches6 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1478  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1479  | 
inductive detect_switches7 :: "('a => bool) => ('b => bool) => ('a * 'b list) => bool"
 | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1480  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1481  | 
"detect_switches7 P Q (a, [])"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1482  | 
| "P a ==> Q x ==> detect_switches7 P Q (a, x#xs)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1483  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1484  | 
code_pred [skip_proof] detect_switches7 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1485  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1486  | 
thm detect_switches7.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1487  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1488  | 
inductive detect_switches8 :: "nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1489  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1490  | 
"detect_switches8 0"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1491  | 
| "x mod 2 = 0 ==> detect_switches8 (Suc x)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1492  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1493  | 
code_pred [detect_switches, skip_proof] detect_switches8 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1494  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1495  | 
thm detect_switches8.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1496  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1497  | 
inductive detect_switches9 :: "nat => nat => bool"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1498  | 
where  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1499  | 
"detect_switches9 0 0"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1500  | 
| "detect_switches9 0 (Suc x)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1501  | 
| "detect_switches9 (Suc x) 0"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1502  | 
| "x = y ==> detect_switches9 (Suc x) (Suc y)"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1503  | 
| "c1 = c2 ==> detect_switches9 c1 c2"  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1504  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1505  | 
code_pred [detect_switches, skip_proof] detect_switches9 .  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1506  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1507  | 
thm detect_switches9.equation  | 
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1508  | 
|
| 
39762
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1509  | 
text {* The higher-order predicate r is in an output term *}
 | 
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1510  | 
|
| 58310 | 1511  | 
datatype result = Result bool  | 
| 
39762
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1512  | 
|
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1513  | 
inductive fixed_relation :: "'a => bool"  | 
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1514  | 
|
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1515  | 
inductive test_relation_in_output_terms :: "('a => bool) => 'a => result => bool"
 | 
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1516  | 
where  | 
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1517  | 
"test_relation_in_output_terms r x (Result (r x))"  | 
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1518  | 
| "test_relation_in_output_terms r x (Result (fixed_relation x))"  | 
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1519  | 
|
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1520  | 
code_pred test_relation_in_output_terms .  | 
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1521  | 
|
| 
 
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
 
bulwahn 
parents: 
39657 
diff
changeset
 | 
1522  | 
thm test_relation_in_output_terms.equation  | 
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1523  | 
|
| 
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1524  | 
|
| 
39765
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1525  | 
text {*
 | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1526  | 
We want that the argument r is not treated as a higher-order relation, but simply as input.  | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1527  | 
*}  | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1528  | 
|
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1529  | 
inductive test_uninterpreted_relation :: "('a => bool) => 'a list => bool"
 | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1530  | 
where  | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1531  | 
"list_all r xs ==> test_uninterpreted_relation r xs"  | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1532  | 
|
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1533  | 
code_pred (modes: i => i => bool) test_uninterpreted_relation .  | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1534  | 
|
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1535  | 
thm test_uninterpreted_relation.equation  | 
| 
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1536  | 
|
| 
39786
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1537  | 
inductive list_ex'  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1538  | 
where  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1539  | 
"P x ==> list_ex' P (x#xs)"  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1540  | 
| "list_ex' P xs ==> list_ex' P (x#xs)"  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1541  | 
|
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1542  | 
code_pred list_ex' .  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1543  | 
|
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1544  | 
inductive test_uninterpreted_relation2 :: "('a => bool) => 'a list => bool"
 | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1545  | 
where  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1546  | 
"list_ex r xs ==> test_uninterpreted_relation2 r xs"  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1547  | 
| "list_ex' r xs ==> test_uninterpreted_relation2 r xs"  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1548  | 
|
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1549  | 
text {* Proof procedure cannot handle this situation yet. *}
 | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1550  | 
|
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1551  | 
code_pred (modes: i => i => bool) [skip_proof] test_uninterpreted_relation2 .  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1552  | 
|
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1553  | 
thm test_uninterpreted_relation2.equation  | 
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1554  | 
|
| 
 
30c077288dfe
added test case for predicate arguments in higher-order argument position
 
bulwahn 
parents: 
39784 
diff
changeset
 | 
1555  | 
|
| 39784 | 1556  | 
text {* Trivial predicate *}
 | 
1557  | 
||
1558  | 
inductive implies_itself :: "'a => bool"  | 
|
1559  | 
where  | 
|
1560  | 
"implies_itself x ==> implies_itself x"  | 
|
1561  | 
||
1562  | 
code_pred implies_itself .  | 
|
| 
39765
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1563  | 
|
| 39803 | 1564  | 
text {* Case expressions *}
 | 
1565  | 
||
1566  | 
definition  | 
|
| 55932 | 1567  | 
"map_prods xs ys = (map (%((a, b), c). (a, b, c)) xs = ys)"  | 
| 39803 | 1568  | 
|
| 55932 | 1569  | 
code_pred [inductify] map_prods .  | 
| 
39765
 
19cb8d558166
adding test case for interpretation of arguments that are predicates simply as input
 
bulwahn 
parents: 
39762 
diff
changeset
 | 
1570  | 
|
| 
39655
 
8ad7fe9d6f0b
splitting Predicate_Compile_Examples into Examples and Predicate_Compile_Tests
 
bulwahn 
parents:  
diff
changeset
 | 
1571  | 
end  |