| author | haftmann |
| Wed, 03 Mar 2010 17:21:45 +0100 | |
| changeset 35550 | e2bc7f8d8d51 |
| parent 35324 | c9f428269b38 |
| child 35884 | 362bfc2ca0ee |
| permissions | -rw-r--r-- |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
1 |
theory Predicate_Compile_Alternative_Defs |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
2 |
imports "../Predicate_Compile" |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
3 |
begin |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
4 |
|
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
5 |
section {* Common constants *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
6 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
7 |
declare HOL.if_bool_eq_disj[code_pred_inline] |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
8 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
9 |
setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
10 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
11 |
section {* Pairs *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
12 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
13 |
setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name split}] *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
14 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
15 |
section {* Bounded quantifiers *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
16 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
17 |
declare Ball_def[code_pred_inline] |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
18 |
declare Bex_def[code_pred_inline] |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
19 |
|
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
20 |
section {* Set operations *}
|
|
33623
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents:
33326
diff
changeset
|
21 |
|
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
22 |
declare Collect_def[code_pred_inline] |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
23 |
declare mem_def[code_pred_inline] |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
24 |
|
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
25 |
declare eq_reflection[OF empty_def, code_pred_inline] |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
26 |
declare insert_code[code_pred_def] |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
27 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
28 |
declare subset_iff[code_pred_inline] |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
29 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
30 |
declare Int_def[code_pred_inline] |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
31 |
declare eq_reflection[OF Un_def, code_pred_inline] |
|
33623
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents:
33326
diff
changeset
|
32 |
declare eq_reflection[OF UNION_def, code_pred_inline] |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
33 |
|
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
34 |
lemma Diff[code_pred_inline]: |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
35 |
"(A - B) = (%x. A x \<and> \<not> B x)" |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
36 |
by (auto simp add: mem_def) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
37 |
|
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
38 |
lemma set_equality[code_pred_inline]: |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
39 |
"(A = B) = ((\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x))" |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
40 |
by (fastsimp simp add: mem_def) |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
41 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
42 |
section {* Setup for Numerals *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
43 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
44 |
setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
45 |
setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
46 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
47 |
setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *}
|
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
48 |
|
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
49 |
section {* Alternative list definitions *}
|
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
50 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
51 |
text {* size simps are not yet added to the Spec_Rules interface. So they are just added manually here! *}
|
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
52 |
|
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
53 |
lemma [code_pred_def]: |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
54 |
"length [] = 0" |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
55 |
"length (x # xs) = Suc (length xs)" |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
56 |
by auto |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
57 |
|
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
58 |
subsection {* Alternative rules for set *}
|
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
59 |
|
| 33628 | 60 |
lemma set_ConsI1 [code_pred_intro]: |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
61 |
"set (x # xs) x" |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
62 |
unfolding mem_def[symmetric, of _ x] |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
63 |
by auto |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
64 |
|
| 33628 | 65 |
lemma set_ConsI2 [code_pred_intro]: |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
66 |
"set xs x ==> set (x' # xs) x" |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
67 |
unfolding mem_def[symmetric, of _ x] |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
68 |
by auto |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
69 |
|
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
70 |
code_pred [skip_proof] set |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
71 |
proof - |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
72 |
case set |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
73 |
from this show thesis |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
74 |
apply (case_tac xb) |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
75 |
apply auto |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
76 |
unfolding mem_def[symmetric, of _ xc] |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
77 |
apply auto |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
78 |
unfolding mem_def |
|
33326
7d0288d90535
improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents:
33250
diff
changeset
|
79 |
apply fastsimp |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
80 |
done |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
81 |
qed |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
82 |
|
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
83 |
subsection {* Alternative rules for list_all2 *}
|
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
84 |
|
| 33628 | 85 |
lemma list_all2_NilI [code_pred_intro]: "list_all2 P [] []" |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
86 |
by auto |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
87 |
|
| 33628 | 88 |
lemma list_all2_ConsI [code_pred_intro]: "list_all2 P xs ys ==> P x y ==> list_all2 P (x#xs) (y#ys)" |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
89 |
by auto |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
90 |
|
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
91 |
code_pred [skip_proof] list_all2 |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
92 |
proof - |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
93 |
case list_all2 |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
94 |
from this show thesis |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
95 |
apply - |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
96 |
apply (case_tac xa) |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
97 |
apply (case_tac xb) |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
98 |
apply auto |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33628
diff
changeset
|
99 |
apply (case_tac xb) |
|
33150
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
100 |
apply auto |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
101 |
done |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
102 |
qed |
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
103 |
|
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
104 |
|
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
105 |
|
|
75eddea4abef
further changes due to the previous merge in the predicate compiler
bulwahn
parents:
diff
changeset
|
106 |
end |