author | haftmann |
Wed, 30 Sep 2009 17:04:21 +0200 | |
changeset 32779 | 371c7f74282d |
parent 32705 | 04ce6bb14d85 |
child 32782 | faf347097852 |
permissions | -rw-r--r-- |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
1 |
(* Title: HOL/Predicate.thy |
30328 | 2 |
Author: Stefan Berghofer and Lukas Bulwahn and Florian Haftmann, TU Muenchen |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
3 |
*) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
4 |
|
30328 | 5 |
header {* Predicates as relations and enumerations *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
6 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
7 |
theory Predicate |
23708 | 8 |
imports Inductive Relation |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
9 |
begin |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
10 |
|
30328 | 11 |
notation |
12 |
inf (infixl "\<sqinter>" 70) and |
|
13 |
sup (infixl "\<squnion>" 65) and |
|
14 |
Inf ("\<Sqinter>_" [900] 900) and |
|
15 |
Sup ("\<Squnion>_" [900] 900) and |
|
16 |
top ("\<top>") and |
|
17 |
bot ("\<bottom>") |
|
18 |
||
19 |
||
20 |
subsection {* Predicates as (complete) lattices *} |
|
21 |
||
22 |
subsubsection {* @{const sup} on @{typ bool} *} |
|
23 |
||
24 |
lemma sup_boolI1: |
|
25 |
"P \<Longrightarrow> P \<squnion> Q" |
|
26 |
by (simp add: sup_bool_eq) |
|
27 |
||
28 |
lemma sup_boolI2: |
|
29 |
"Q \<Longrightarrow> P \<squnion> Q" |
|
30 |
by (simp add: sup_bool_eq) |
|
31 |
||
32 |
lemma sup_boolE: |
|
33 |
"P \<squnion> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R" |
|
34 |
by (auto simp add: sup_bool_eq) |
|
35 |
||
36 |
||
32779 | 37 |
subsubsection {* Equality *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
38 |
|
26797
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
39 |
lemma pred_equals_eq: "((\<lambda>x. x \<in> R) = (\<lambda>x. x \<in> S)) = (R = S)" |
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
40 |
by (simp add: mem_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
41 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
42 |
lemma pred_equals_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) = (\<lambda>x y. (x, y) \<in> S)) = (R = S)" |
26797
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
43 |
by (simp add: expand_fun_eq mem_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
44 |
|
32779 | 45 |
|
46 |
subsubsection {* Order relation *} |
|
47 |
||
26797
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
48 |
lemma pred_subset_eq: "((\<lambda>x. x \<in> R) <= (\<lambda>x. x \<in> S)) = (R <= S)" |
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
49 |
by (simp add: mem_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
50 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
51 |
lemma pred_subset_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) <= (\<lambda>x y. (x, y) \<in> S)) = (R <= S)" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
52 |
by fast |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
53 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
54 |
|
30328 | 55 |
subsubsection {* Top and bottom elements *} |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
56 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
57 |
lemma top1I [intro!]: "top x" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
58 |
by (simp add: top_fun_eq top_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
59 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
60 |
lemma top2I [intro!]: "top x y" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
61 |
by (simp add: top_fun_eq top_bool_eq) |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
62 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
63 |
lemma bot1E [elim!]: "bot x \<Longrightarrow> P" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
64 |
by (simp add: bot_fun_eq bot_bool_eq) |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
65 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
66 |
lemma bot2E [elim!]: "bot x y \<Longrightarrow> P" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
67 |
by (simp add: bot_fun_eq bot_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
68 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
69 |
lemma bot_empty_eq: "bot = (\<lambda>x. x \<in> {})" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
70 |
by (auto simp add: expand_fun_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
71 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
72 |
lemma bot_empty_eq2: "bot = (\<lambda>x y. (x, y) \<in> {})" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
73 |
by (auto simp add: expand_fun_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
74 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
75 |
|
30328 | 76 |
subsubsection {* Binary union *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
77 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
78 |
lemma sup1_iff: "sup A B x \<longleftrightarrow> A x | B x" |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
79 |
by (simp add: sup_fun_eq sup_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
80 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
81 |
lemma sup2_iff: "sup A B x y \<longleftrightarrow> A x y | B x y" |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
82 |
by (simp add: sup_fun_eq sup_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
83 |
|
32703
7f9e05b3d0fb
removed potentially dangerous rules from pred_set_conv
haftmann
parents:
32601
diff
changeset
|
84 |
lemma sup_Un_eq: "sup (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<union> S)" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
85 |
by (simp add: sup1_iff expand_fun_eq) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
86 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
87 |
lemma sup_Un_eq2 [pred_set_conv]: "sup (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<union> S)" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
88 |
by (simp add: sup2_iff expand_fun_eq) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
89 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
90 |
lemma sup1I1 [elim?]: "A x \<Longrightarrow> sup A B x" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
91 |
by (simp add: sup1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
92 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
93 |
lemma sup2I1 [elim?]: "A x y \<Longrightarrow> sup A B x y" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
94 |
by (simp add: sup2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
95 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
96 |
lemma sup1I2 [elim?]: "B x \<Longrightarrow> sup A B x" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
97 |
by (simp add: sup1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
98 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
99 |
lemma sup2I2 [elim?]: "B x y \<Longrightarrow> sup A B x y" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
100 |
by (simp add: sup2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
101 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
102 |
text {* |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
103 |
\medskip Classical introduction rule: no commitment to @{text A} vs |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
104 |
@{text B}. |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
105 |
*} |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
106 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
107 |
lemma sup1CI [intro!]: "(~ B x ==> A x) ==> sup A B x" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
108 |
by (auto simp add: sup1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
109 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
110 |
lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
111 |
by (auto simp add: sup2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
112 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
113 |
lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
114 |
by (simp add: sup1_iff) iprover |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
115 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
116 |
lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
117 |
by (simp add: sup2_iff) iprover |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
118 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
119 |
|
30328 | 120 |
subsubsection {* Binary intersection *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
121 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
122 |
lemma inf1_iff: "inf A B x \<longleftrightarrow> A x \<and> B x" |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
123 |
by (simp add: inf_fun_eq inf_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
124 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
125 |
lemma inf2_iff: "inf A B x y \<longleftrightarrow> A x y \<and> B x y" |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
126 |
by (simp add: inf_fun_eq inf_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
127 |
|
32703
7f9e05b3d0fb
removed potentially dangerous rules from pred_set_conv
haftmann
parents:
32601
diff
changeset
|
128 |
lemma inf_Int_eq: "inf (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<inter> S)" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
129 |
by (simp add: inf1_iff expand_fun_eq) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
130 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
131 |
lemma inf_Int_eq2 [pred_set_conv]: "inf (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<inter> S)" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
132 |
by (simp add: inf2_iff expand_fun_eq) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
133 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
134 |
lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
135 |
by (simp add: inf1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
136 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
137 |
lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
138 |
by (simp add: inf2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
139 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
140 |
lemma inf1D1: "inf A B x ==> A x" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
141 |
by (simp add: inf1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
142 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
143 |
lemma inf2D1: "inf A B x y ==> A x y" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
144 |
by (simp add: inf2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
145 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
146 |
lemma inf1D2: "inf A B x ==> B x" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
147 |
by (simp add: inf1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
148 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
149 |
lemma inf2D2: "inf A B x y ==> B x y" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
150 |
by (simp add: inf2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
151 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
152 |
lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
153 |
by (simp add: inf1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
154 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
155 |
lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
156 |
by (simp add: inf2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
157 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
158 |
|
30328 | 159 |
subsubsection {* Unions of families *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
160 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
161 |
lemma SUP1_iff: "(SUP x:A. B x) b = (EX x:A. B x b)" |
24345 | 162 |
by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
163 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
164 |
lemma SUP2_iff: "(SUP x:A. B x) b c = (EX x:A. B x b c)" |
24345 | 165 |
by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
166 |
|
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
167 |
lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
168 |
by (auto simp add: SUP1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
169 |
|
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
170 |
lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
171 |
by (auto simp add: SUP2_iff) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
172 |
|
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
173 |
lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
174 |
by (auto simp add: SUP1_iff) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
175 |
|
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
176 |
lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
177 |
by (auto simp add: SUP2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
178 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
179 |
lemma SUP_UN_eq: "(SUP i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (UN i. r i))" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
180 |
by (simp add: SUP1_iff expand_fun_eq) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
181 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
182 |
lemma SUP_UN_eq2: "(SUP i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (UN i. r i))" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
183 |
by (simp add: SUP2_iff expand_fun_eq) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
184 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
185 |
|
30328 | 186 |
subsubsection {* Intersections of families *} |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
187 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
188 |
lemma INF1_iff: "(INF x:A. B x) b = (ALL x:A. B x b)" |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
189 |
by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast |
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
190 |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
191 |
lemma INF2_iff: "(INF x:A. B x) b c = (ALL x:A. B x b c)" |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
192 |
by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast |
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
193 |
|
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
194 |
lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
195 |
by (auto simp add: INF1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
196 |
|
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
197 |
lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
198 |
by (auto simp add: INF2_iff) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
199 |
|
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
200 |
lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
201 |
by (auto simp add: INF1_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
202 |
|
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
203 |
lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
204 |
by (auto simp add: INF2_iff) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
205 |
|
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
206 |
lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
207 |
by (auto simp add: INF1_iff) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
208 |
|
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
209 |
lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
210 |
by (auto simp add: INF2_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
211 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
212 |
lemma INF_INT_eq: "(INF i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (INT i. r i))" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
213 |
by (simp add: INF1_iff expand_fun_eq) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
214 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
215 |
lemma INF_INT_eq2: "(INF i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (INT i. r i))" |
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
216 |
by (simp add: INF2_iff expand_fun_eq) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
217 |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
218 |
|
30328 | 219 |
subsection {* Predicates as relations *} |
220 |
||
221 |
subsubsection {* Composition *} |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
222 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
223 |
inductive |
32235
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
krauss
parents:
31932
diff
changeset
|
224 |
pred_comp :: "['a => 'b => bool, 'b => 'c => bool] => 'a => 'c => bool" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
225 |
(infixr "OO" 75) |
32235
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
krauss
parents:
31932
diff
changeset
|
226 |
for r :: "'a => 'b => bool" and s :: "'b => 'c => bool" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
227 |
where |
32235
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
krauss
parents:
31932
diff
changeset
|
228 |
pred_compI [intro]: "r a b ==> s b c ==> (r OO s) a c" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
229 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
230 |
inductive_cases pred_compE [elim!]: "(r OO s) a c" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
231 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
232 |
lemma pred_comp_rel_comp_eq [pred_set_conv]: |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
233 |
"((\<lambda>x y. (x, y) \<in> r) OO (\<lambda>x y. (x, y) \<in> s)) = (\<lambda>x y. (x, y) \<in> r O s)" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
234 |
by (auto simp add: expand_fun_eq elim: pred_compE) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
235 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
236 |
|
30328 | 237 |
subsubsection {* Converse *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
238 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
239 |
inductive |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
240 |
conversep :: "('a => 'b => bool) => 'b => 'a => bool" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
241 |
("(_^--1)" [1000] 1000) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
242 |
for r :: "'a => 'b => bool" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
243 |
where |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
244 |
conversepI: "r a b ==> r^--1 b a" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
245 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
246 |
notation (xsymbols) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
247 |
conversep ("(_\<inverse>\<inverse>)" [1000] 1000) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
248 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
249 |
lemma conversepD: |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
250 |
assumes ab: "r^--1 a b" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
251 |
shows "r b a" using ab |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
252 |
by cases simp |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
253 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
254 |
lemma conversep_iff [iff]: "r^--1 a b = r b a" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
255 |
by (iprover intro: conversepI dest: conversepD) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
256 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
257 |
lemma conversep_converse_eq [pred_set_conv]: |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
258 |
"(\<lambda>x y. (x, y) \<in> r)^--1 = (\<lambda>x y. (x, y) \<in> r^-1)" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
259 |
by (auto simp add: expand_fun_eq) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
260 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
261 |
lemma conversep_conversep [simp]: "(r^--1)^--1 = r" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
262 |
by (iprover intro: order_antisym conversepI dest: conversepD) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
263 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
264 |
lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
265 |
by (iprover intro: order_antisym conversepI pred_compI |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
266 |
elim: pred_compE dest: conversepD) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
267 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
268 |
lemma converse_meet: "(inf r s)^--1 = inf r^--1 s^--1" |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
269 |
by (simp add: inf_fun_eq inf_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
270 |
(iprover intro: conversepI ext dest: conversepD) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
271 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
272 |
lemma converse_join: "(sup r s)^--1 = sup r^--1 s^--1" |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22259
diff
changeset
|
273 |
by (simp add: sup_fun_eq sup_bool_eq) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
274 |
(iprover intro: conversepI ext dest: conversepD) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
275 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
276 |
lemma conversep_noteq [simp]: "(op ~=)^--1 = op ~=" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
277 |
by (auto simp add: expand_fun_eq) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
278 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
279 |
lemma conversep_eq [simp]: "(op =)^--1 = op =" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
280 |
by (auto simp add: expand_fun_eq) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
281 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
282 |
|
30328 | 283 |
subsubsection {* Domain *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
284 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
285 |
inductive |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
286 |
DomainP :: "('a => 'b => bool) => 'a => bool" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
287 |
for r :: "'a => 'b => bool" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
288 |
where |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
289 |
DomainPI [intro]: "r a b ==> DomainP r a" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
290 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
291 |
inductive_cases DomainPE [elim!]: "DomainP r a" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
292 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
293 |
lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Domain r)" |
26797
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
294 |
by (blast intro!: Orderings.order_antisym predicate1I) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
295 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
296 |
|
30328 | 297 |
subsubsection {* Range *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
298 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
299 |
inductive |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
300 |
RangeP :: "('a => 'b => bool) => 'b => bool" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
301 |
for r :: "'a => 'b => bool" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
302 |
where |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
303 |
RangePI [intro]: "r a b ==> RangeP r b" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
304 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
305 |
inductive_cases RangePE [elim!]: "RangeP r b" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
306 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
307 |
lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Range r)" |
26797
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
308 |
by (blast intro!: Orderings.order_antisym predicate1I) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
309 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
310 |
|
30328 | 311 |
subsubsection {* Inverse image *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
312 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
313 |
definition |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
314 |
inv_imagep :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" where |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
315 |
"inv_imagep r f == %x y. r (f x) (f y)" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
316 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
317 |
lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
318 |
by (simp add: inv_image_def inv_imagep_def) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
319 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
320 |
lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
321 |
by (simp add: inv_imagep_def) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
322 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
323 |
|
30328 | 324 |
subsubsection {* Powerset *} |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
325 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
326 |
definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
327 |
"Powp A == \<lambda>B. \<forall>x \<in> B. A x" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
328 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
329 |
lemma Powp_Pow_eq [pred_set_conv]: "Powp (\<lambda>x. x \<in> A) = (\<lambda>x. x \<in> Pow A)" |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
330 |
by (auto simp add: Powp_def expand_fun_eq) |
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
331 |
|
26797
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
332 |
lemmas Powp_mono [mono] = Pow_mono [to_pred pred_subset_eq] |
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
333 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
334 |
|
30328 | 335 |
subsubsection {* Properties of relations *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
336 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
337 |
abbreviation antisymP :: "('a => 'a => bool) => bool" where |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
338 |
"antisymP r == antisym {(x, y). r x y}" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
339 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
340 |
abbreviation transP :: "('a => 'a => bool) => bool" where |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
341 |
"transP r == trans {(x, y). r x y}" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
342 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
343 |
abbreviation single_valuedP :: "('a => 'b => bool) => bool" where |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
344 |
"single_valuedP r == single_valued {(x, y). r x y}" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
345 |
|
30328 | 346 |
|
347 |
subsection {* Predicates as enumerations *} |
|
348 |
||
349 |
subsubsection {* The type of predicate enumerations (a monad) *} |
|
350 |
||
351 |
datatype 'a pred = Pred "'a \<Rightarrow> bool" |
|
352 |
||
353 |
primrec eval :: "'a pred \<Rightarrow> 'a \<Rightarrow> bool" where |
|
354 |
eval_pred: "eval (Pred f) = f" |
|
355 |
||
356 |
lemma Pred_eval [simp]: |
|
357 |
"Pred (eval x) = x" |
|
358 |
by (cases x) simp |
|
359 |
||
360 |
lemma eval_inject: "eval x = eval y \<longleftrightarrow> x = y" |
|
361 |
by (cases x) auto |
|
362 |
||
363 |
definition single :: "'a \<Rightarrow> 'a pred" where |
|
364 |
"single x = Pred ((op =) x)" |
|
365 |
||
366 |
definition bind :: "'a pred \<Rightarrow> ('a \<Rightarrow> 'b pred) \<Rightarrow> 'b pred" (infixl "\<guillemotright>=" 70) where |
|
367 |
"P \<guillemotright>= f = Pred (\<lambda>x. (\<exists>y. eval P y \<and> eval (f y) x))" |
|
368 |
||
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
369 |
instantiation pred :: (type) "{complete_lattice, boolean_algebra}" |
30328 | 370 |
begin |
371 |
||
372 |
definition |
|
373 |
"P \<le> Q \<longleftrightarrow> eval P \<le> eval Q" |
|
374 |
||
375 |
definition |
|
376 |
"P < Q \<longleftrightarrow> eval P < eval Q" |
|
377 |
||
378 |
definition |
|
379 |
"\<bottom> = Pred \<bottom>" |
|
380 |
||
381 |
definition |
|
382 |
"\<top> = Pred \<top>" |
|
383 |
||
384 |
definition |
|
385 |
"P \<sqinter> Q = Pred (eval P \<sqinter> eval Q)" |
|
386 |
||
387 |
definition |
|
388 |
"P \<squnion> Q = Pred (eval P \<squnion> eval Q)" |
|
389 |
||
390 |
definition |
|
31932 | 391 |
[code del]: "\<Sqinter>A = Pred (INFI A eval)" |
30328 | 392 |
|
393 |
definition |
|
31932 | 394 |
[code del]: "\<Squnion>A = Pred (SUPR A eval)" |
30328 | 395 |
|
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
396 |
definition |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
397 |
"- P = Pred (- eval P)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
398 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
399 |
definition |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
400 |
"P - Q = Pred (eval P - eval Q)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
401 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
402 |
instance proof |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
403 |
qed (auto simp add: less_eq_pred_def less_pred_def |
30328 | 404 |
inf_pred_def sup_pred_def bot_pred_def top_pred_def |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
405 |
Inf_pred_def Sup_pred_def uminus_pred_def minus_pred_def fun_Compl_def bool_Compl_def, |
30328 | 406 |
auto simp add: le_fun_def less_fun_def le_bool_def less_bool_def |
407 |
eval_inject mem_def) |
|
408 |
||
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
409 |
end |
30328 | 410 |
|
411 |
lemma bind_bind: |
|
412 |
"(P \<guillemotright>= Q) \<guillemotright>= R = P \<guillemotright>= (\<lambda>x. Q x \<guillemotright>= R)" |
|
413 |
by (auto simp add: bind_def expand_fun_eq) |
|
414 |
||
415 |
lemma bind_single: |
|
416 |
"P \<guillemotright>= single = P" |
|
417 |
by (simp add: bind_def single_def) |
|
418 |
||
419 |
lemma single_bind: |
|
420 |
"single x \<guillemotright>= P = P x" |
|
421 |
by (simp add: bind_def single_def) |
|
422 |
||
423 |
lemma bottom_bind: |
|
424 |
"\<bottom> \<guillemotright>= P = \<bottom>" |
|
425 |
by (auto simp add: bot_pred_def bind_def expand_fun_eq) |
|
426 |
||
427 |
lemma sup_bind: |
|
428 |
"(P \<squnion> Q) \<guillemotright>= R = P \<guillemotright>= R \<squnion> Q \<guillemotright>= R" |
|
429 |
by (auto simp add: bind_def sup_pred_def expand_fun_eq) |
|
430 |
||
431 |
lemma Sup_bind: "(\<Squnion>A \<guillemotright>= f) = \<Squnion>((\<lambda>x. x \<guillemotright>= f) ` A)" |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
432 |
by (auto simp add: bind_def Sup_pred_def SUP1_iff expand_fun_eq) |
30328 | 433 |
|
434 |
lemma pred_iffI: |
|
435 |
assumes "\<And>x. eval A x \<Longrightarrow> eval B x" |
|
436 |
and "\<And>x. eval B x \<Longrightarrow> eval A x" |
|
437 |
shows "A = B" |
|
438 |
proof - |
|
439 |
from assms have "\<And>x. eval A x \<longleftrightarrow> eval B x" by blast |
|
440 |
then show ?thesis by (cases A, cases B) (simp add: expand_fun_eq) |
|
441 |
qed |
|
442 |
||
443 |
lemma singleI: "eval (single x) x" |
|
444 |
unfolding single_def by simp |
|
445 |
||
446 |
lemma singleI_unit: "eval (single ()) x" |
|
447 |
by simp (rule singleI) |
|
448 |
||
449 |
lemma singleE: "eval (single x) y \<Longrightarrow> (y = x \<Longrightarrow> P) \<Longrightarrow> P" |
|
450 |
unfolding single_def by simp |
|
451 |
||
452 |
lemma singleE': "eval (single x) y \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> P" |
|
453 |
by (erule singleE) simp |
|
454 |
||
455 |
lemma bindI: "eval P x \<Longrightarrow> eval (Q x) y \<Longrightarrow> eval (P \<guillemotright>= Q) y" |
|
456 |
unfolding bind_def by auto |
|
457 |
||
458 |
lemma bindE: "eval (R \<guillemotright>= Q) y \<Longrightarrow> (\<And>x. eval R x \<Longrightarrow> eval (Q x) y \<Longrightarrow> P) \<Longrightarrow> P" |
|
459 |
unfolding bind_def by auto |
|
460 |
||
461 |
lemma botE: "eval \<bottom> x \<Longrightarrow> P" |
|
462 |
unfolding bot_pred_def by auto |
|
463 |
||
464 |
lemma supI1: "eval A x \<Longrightarrow> eval (A \<squnion> B) x" |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
465 |
unfolding sup_pred_def by (simp add: sup1_iff) |
30328 | 466 |
|
467 |
lemma supI2: "eval B x \<Longrightarrow> eval (A \<squnion> B) x" |
|
32601
47d0c967c64e
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
haftmann
parents:
32582
diff
changeset
|
468 |
unfolding sup_pred_def by (simp add: sup1_iff) |
30328 | 469 |
|
470 |
lemma supE: "eval (A \<squnion> B) x \<Longrightarrow> (eval A x \<Longrightarrow> P) \<Longrightarrow> (eval B x \<Longrightarrow> P) \<Longrightarrow> P" |
|
471 |
unfolding sup_pred_def by auto |
|
472 |
||
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
473 |
lemma single_not_bot [simp]: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
474 |
"single x \<noteq> \<bottom>" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
475 |
by (auto simp add: single_def bot_pred_def expand_fun_eq) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
476 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
477 |
lemma not_bot: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
478 |
assumes "A \<noteq> \<bottom>" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
479 |
obtains x where "eval A x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
480 |
using assms by (cases A) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
481 |
(auto simp add: bot_pred_def, auto simp add: mem_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
482 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
483 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
484 |
subsubsection {* Emptiness check and definite choice *} |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
485 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
486 |
definition is_empty :: "'a pred \<Rightarrow> bool" where |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
487 |
"is_empty A \<longleftrightarrow> A = \<bottom>" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
488 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
489 |
lemma is_empty_bot: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
490 |
"is_empty \<bottom>" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
491 |
by (simp add: is_empty_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
492 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
493 |
lemma not_is_empty_single: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
494 |
"\<not> is_empty (single x)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
495 |
by (auto simp add: is_empty_def single_def bot_pred_def expand_fun_eq) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
496 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
497 |
lemma is_empty_sup: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
498 |
"is_empty (A \<squnion> B) \<longleftrightarrow> is_empty A \<and> is_empty B" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
499 |
by (auto simp add: is_empty_def intro: sup_eq_bot_eq1 sup_eq_bot_eq2) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
500 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
501 |
definition singleton :: "'a pred \<Rightarrow> 'a" where |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
502 |
"singleton A = (if \<exists>!x. eval A x then THE x. eval A x else undefined)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
503 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
504 |
lemma singleton_eqI: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
505 |
"\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton A = x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
506 |
by (auto simp add: singleton_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
507 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
508 |
lemma eval_singletonI: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
509 |
"\<exists>!x. eval A x \<Longrightarrow> eval A (singleton A)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
510 |
proof - |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
511 |
assume assm: "\<exists>!x. eval A x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
512 |
then obtain x where "eval A x" .. |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
513 |
moreover with assm have "singleton A = x" by (rule singleton_eqI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
514 |
ultimately show ?thesis by simp |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
515 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
516 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
517 |
lemma single_singleton: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
518 |
"\<exists>!x. eval A x \<Longrightarrow> single (singleton A) = A" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
519 |
proof - |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
520 |
assume assm: "\<exists>!x. eval A x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
521 |
then have "eval A (singleton A)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
522 |
by (rule eval_singletonI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
523 |
moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton A = x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
524 |
by (rule singleton_eqI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
525 |
ultimately have "eval (single (singleton A)) = eval A" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
526 |
by (simp (no_asm_use) add: single_def expand_fun_eq) blast |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
527 |
then show ?thesis by (simp add: eval_inject) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
528 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
529 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
530 |
lemma singleton_undefinedI: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
531 |
"\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton A = undefined" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
532 |
by (simp add: singleton_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
533 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
534 |
lemma singleton_bot: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
535 |
"singleton \<bottom> = undefined" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
536 |
by (auto simp add: bot_pred_def intro: singleton_undefinedI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
537 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
538 |
lemma singleton_single: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
539 |
"singleton (single x) = x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
540 |
by (auto simp add: intro: singleton_eqI singleI elim: singleE) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
541 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
542 |
lemma singleton_sup_single_single: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
543 |
"singleton (single x \<squnion> single y) = (if x = y then x else undefined)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
544 |
proof (cases "x = y") |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
545 |
case True then show ?thesis by (simp add: singleton_single) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
546 |
next |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
547 |
case False |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
548 |
have "eval (single x \<squnion> single y) x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
549 |
and "eval (single x \<squnion> single y) y" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
550 |
by (auto intro: supI1 supI2 singleI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
551 |
with False have "\<not> (\<exists>!z. eval (single x \<squnion> single y) z)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
552 |
by blast |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
553 |
then have "singleton (single x \<squnion> single y) = undefined" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
554 |
by (rule singleton_undefinedI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
555 |
with False show ?thesis by simp |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
556 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
557 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
558 |
lemma singleton_sup_aux: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
559 |
"singleton (A \<squnion> B) = (if A = \<bottom> then singleton B |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
560 |
else if B = \<bottom> then singleton A |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
561 |
else singleton |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
562 |
(single (singleton A) \<squnion> single (singleton B)))" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
563 |
proof (cases "(\<exists>!x. eval A x) \<and> (\<exists>!y. eval B y)") |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
564 |
case True then show ?thesis by (simp add: single_singleton) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
565 |
next |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
566 |
case False |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
567 |
from False have A_or_B: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
568 |
"singleton A = undefined \<or> singleton B = undefined" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
569 |
by (auto intro!: singleton_undefinedI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
570 |
then have rhs: "singleton |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
571 |
(single (singleton A) \<squnion> single (singleton B)) = undefined" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
572 |
by (auto simp add: singleton_sup_single_single singleton_single) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
573 |
from False have not_unique: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
574 |
"\<not> (\<exists>!x. eval A x) \<or> \<not> (\<exists>!y. eval B y)" by simp |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
575 |
show ?thesis proof (cases "A \<noteq> \<bottom> \<and> B \<noteq> \<bottom>") |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
576 |
case True |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
577 |
then obtain a b where a: "eval A a" and b: "eval B b" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
578 |
by (blast elim: not_bot) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
579 |
with True not_unique have "\<not> (\<exists>!x. eval (A \<squnion> B) x)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
580 |
by (auto simp add: sup_pred_def bot_pred_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
581 |
then have "singleton (A \<squnion> B) = undefined" by (rule singleton_undefinedI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
582 |
with True rhs show ?thesis by simp |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
583 |
next |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
584 |
case False then show ?thesis by auto |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
585 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
586 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
587 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
588 |
lemma singleton_sup: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
589 |
"singleton (A \<squnion> B) = (if A = \<bottom> then singleton B |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
590 |
else if B = \<bottom> then singleton A |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
591 |
else if singleton A = singleton B then singleton A else undefined)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
592 |
using singleton_sup_aux [of A B] by (simp only: singleton_sup_single_single) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
593 |
|
30328 | 594 |
|
595 |
subsubsection {* Derived operations *} |
|
596 |
||
597 |
definition if_pred :: "bool \<Rightarrow> unit pred" where |
|
598 |
if_pred_eq: "if_pred b = (if b then single () else \<bottom>)" |
|
599 |
||
600 |
definition not_pred :: "unit pred \<Rightarrow> unit pred" where |
|
601 |
not_pred_eq: "not_pred P = (if eval P () then \<bottom> else single ())" |
|
602 |
||
603 |
lemma if_predI: "P \<Longrightarrow> eval (if_pred P) ()" |
|
604 |
unfolding if_pred_eq by (auto intro: singleI) |
|
605 |
||
606 |
lemma if_predE: "eval (if_pred b) x \<Longrightarrow> (b \<Longrightarrow> x = () \<Longrightarrow> P) \<Longrightarrow> P" |
|
607 |
unfolding if_pred_eq by (cases b) (auto elim: botE) |
|
608 |
||
609 |
lemma not_predI: "\<not> P \<Longrightarrow> eval (not_pred (Pred (\<lambda>u. P))) ()" |
|
610 |
unfolding not_pred_eq eval_pred by (auto intro: singleI) |
|
611 |
||
612 |
lemma not_predI': "\<not> eval P () \<Longrightarrow> eval (not_pred P) ()" |
|
613 |
unfolding not_pred_eq by (auto intro: singleI) |
|
614 |
||
615 |
lemma not_predE: "eval (not_pred (Pred (\<lambda>u. P))) x \<Longrightarrow> (\<not> P \<Longrightarrow> thesis) \<Longrightarrow> thesis" |
|
616 |
unfolding not_pred_eq |
|
617 |
by (auto split: split_if_asm elim: botE) |
|
618 |
||
619 |
lemma not_predE': "eval (not_pred P) x \<Longrightarrow> (\<not> eval P x \<Longrightarrow> thesis) \<Longrightarrow> thesis" |
|
620 |
unfolding not_pred_eq |
|
621 |
by (auto split: split_if_asm elim: botE) |
|
622 |
||
623 |
||
624 |
subsubsection {* Implementation *} |
|
625 |
||
626 |
datatype 'a seq = Empty | Insert "'a" "'a pred" | Join "'a pred" "'a seq" |
|
627 |
||
628 |
primrec pred_of_seq :: "'a seq \<Rightarrow> 'a pred" where |
|
629 |
"pred_of_seq Empty = \<bottom>" |
|
630 |
| "pred_of_seq (Insert x P) = single x \<squnion> P" |
|
631 |
| "pred_of_seq (Join P xq) = P \<squnion> pred_of_seq xq" |
|
632 |
||
633 |
definition Seq :: "(unit \<Rightarrow> 'a seq) \<Rightarrow> 'a pred" where |
|
634 |
"Seq f = pred_of_seq (f ())" |
|
635 |
||
636 |
code_datatype Seq |
|
637 |
||
638 |
primrec member :: "'a seq \<Rightarrow> 'a \<Rightarrow> bool" where |
|
639 |
"member Empty x \<longleftrightarrow> False" |
|
640 |
| "member (Insert y P) x \<longleftrightarrow> x = y \<or> eval P x" |
|
641 |
| "member (Join P xq) x \<longleftrightarrow> eval P x \<or> member xq x" |
|
642 |
||
643 |
lemma eval_member: |
|
644 |
"member xq = eval (pred_of_seq xq)" |
|
645 |
proof (induct xq) |
|
646 |
case Empty show ?case |
|
647 |
by (auto simp add: expand_fun_eq elim: botE) |
|
648 |
next |
|
649 |
case Insert show ?case |
|
650 |
by (auto simp add: expand_fun_eq elim: supE singleE intro: supI1 supI2 singleI) |
|
651 |
next |
|
652 |
case Join then show ?case |
|
653 |
by (auto simp add: expand_fun_eq elim: supE intro: supI1 supI2) |
|
654 |
qed |
|
655 |
||
656 |
lemma eval_code [code]: "eval (Seq f) = member (f ())" |
|
657 |
unfolding Seq_def by (rule sym, rule eval_member) |
|
658 |
||
659 |
lemma single_code [code]: |
|
660 |
"single x = Seq (\<lambda>u. Insert x \<bottom>)" |
|
661 |
unfolding Seq_def by simp |
|
662 |
||
663 |
primrec "apply" :: "('a \<Rightarrow> 'b Predicate.pred) \<Rightarrow> 'a seq \<Rightarrow> 'b seq" where |
|
664 |
"apply f Empty = Empty" |
|
665 |
| "apply f (Insert x P) = Join (f x) (Join (P \<guillemotright>= f) Empty)" |
|
666 |
| "apply f (Join P xq) = Join (P \<guillemotright>= f) (apply f xq)" |
|
667 |
||
668 |
lemma apply_bind: |
|
669 |
"pred_of_seq (apply f xq) = pred_of_seq xq \<guillemotright>= f" |
|
670 |
proof (induct xq) |
|
671 |
case Empty show ?case |
|
672 |
by (simp add: bottom_bind) |
|
673 |
next |
|
674 |
case Insert show ?case |
|
675 |
by (simp add: single_bind sup_bind) |
|
676 |
next |
|
677 |
case Join then show ?case |
|
678 |
by (simp add: sup_bind) |
|
679 |
qed |
|
680 |
||
681 |
lemma bind_code [code]: |
|
682 |
"Seq g \<guillemotright>= f = Seq (\<lambda>u. apply f (g ()))" |
|
683 |
unfolding Seq_def by (rule sym, rule apply_bind) |
|
684 |
||
685 |
lemma bot_set_code [code]: |
|
686 |
"\<bottom> = Seq (\<lambda>u. Empty)" |
|
687 |
unfolding Seq_def by simp |
|
688 |
||
30376 | 689 |
primrec adjunct :: "'a pred \<Rightarrow> 'a seq \<Rightarrow> 'a seq" where |
690 |
"adjunct P Empty = Join P Empty" |
|
691 |
| "adjunct P (Insert x Q) = Insert x (Q \<squnion> P)" |
|
692 |
| "adjunct P (Join Q xq) = Join Q (adjunct P xq)" |
|
693 |
||
694 |
lemma adjunct_sup: |
|
695 |
"pred_of_seq (adjunct P xq) = P \<squnion> pred_of_seq xq" |
|
696 |
by (induct xq) (simp_all add: sup_assoc sup_commute sup_left_commute) |
|
697 |
||
30328 | 698 |
lemma sup_code [code]: |
699 |
"Seq f \<squnion> Seq g = Seq (\<lambda>u. case f () |
|
700 |
of Empty \<Rightarrow> g () |
|
701 |
| Insert x P \<Rightarrow> Insert x (P \<squnion> Seq g) |
|
30376 | 702 |
| Join P xq \<Rightarrow> adjunct (Seq g) (Join P xq))" |
30328 | 703 |
proof (cases "f ()") |
704 |
case Empty |
|
705 |
thus ?thesis |
|
30376 | 706 |
unfolding Seq_def by (simp add: sup_commute [of "\<bottom>"] sup_bot) |
30328 | 707 |
next |
708 |
case Insert |
|
709 |
thus ?thesis |
|
710 |
unfolding Seq_def by (simp add: sup_assoc) |
|
711 |
next |
|
712 |
case Join |
|
713 |
thus ?thesis |
|
30376 | 714 |
unfolding Seq_def |
715 |
by (simp add: adjunct_sup sup_assoc sup_commute sup_left_commute) |
|
30328 | 716 |
qed |
717 |
||
30430
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
718 |
primrec contained :: "'a seq \<Rightarrow> 'a pred \<Rightarrow> bool" where |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
719 |
"contained Empty Q \<longleftrightarrow> True" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
720 |
| "contained (Insert x P) Q \<longleftrightarrow> eval Q x \<and> P \<le> Q" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
721 |
| "contained (Join P xq) Q \<longleftrightarrow> P \<le> Q \<and> contained xq Q" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
722 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
723 |
lemma single_less_eq_eval: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
724 |
"single x \<le> P \<longleftrightarrow> eval P x" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
725 |
by (auto simp add: single_def less_eq_pred_def mem_def) |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
726 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
727 |
lemma contained_less_eq: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
728 |
"contained xq Q \<longleftrightarrow> pred_of_seq xq \<le> Q" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
729 |
by (induct xq) (simp_all add: single_less_eq_eval) |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
730 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
731 |
lemma less_eq_pred_code [code]: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
732 |
"Seq f \<le> Q = (case f () |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
733 |
of Empty \<Rightarrow> True |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
734 |
| Insert x P \<Rightarrow> eval Q x \<and> P \<le> Q |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
735 |
| Join P xq \<Rightarrow> P \<le> Q \<and> contained xq Q)" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
736 |
by (cases "f ()") |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
737 |
(simp_all add: Seq_def single_less_eq_eval contained_less_eq) |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
738 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
739 |
lemma eq_pred_code [code]: |
31133 | 740 |
fixes P Q :: "'a pred" |
30430
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
741 |
shows "eq_class.eq P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
742 |
unfolding eq by auto |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
743 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
744 |
lemma [code]: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
745 |
"pred_case f P = f (eval P)" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
746 |
by (cases P) simp |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
747 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
748 |
lemma [code]: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
749 |
"pred_rec f P = f (eval P)" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
750 |
by (cases P) simp |
30328 | 751 |
|
31105
95f66b234086
added general preprocessing of equality in predicates for code generation
bulwahn
parents:
30430
diff
changeset
|
752 |
inductive eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where "eq x x" |
95f66b234086
added general preprocessing of equality in predicates for code generation
bulwahn
parents:
30430
diff
changeset
|
753 |
|
95f66b234086
added general preprocessing of equality in predicates for code generation
bulwahn
parents:
30430
diff
changeset
|
754 |
lemma eq_is_eq: "eq x y \<equiv> (x = y)" |
31108 | 755 |
by (rule eq_reflection) (auto intro: eq.intros elim: eq.cases) |
30948 | 756 |
|
31216 | 757 |
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pred \<Rightarrow> 'b pred" where |
758 |
"map f P = P \<guillemotright>= (single o f)" |
|
759 |
||
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
760 |
primrec null :: "'a seq \<Rightarrow> bool" where |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
761 |
"null Empty \<longleftrightarrow> True" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
762 |
| "null (Insert x P) \<longleftrightarrow> False" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
763 |
| "null (Join P xq) \<longleftrightarrow> is_empty P \<and> null xq" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
764 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
765 |
lemma null_is_empty: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
766 |
"null xq \<longleftrightarrow> is_empty (pred_of_seq xq)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
767 |
by (induct xq) (simp_all add: is_empty_bot not_is_empty_single is_empty_sup) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
768 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
769 |
lemma is_empty_code [code]: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
770 |
"is_empty (Seq f) \<longleftrightarrow> null (f ())" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
771 |
by (simp add: null_is_empty Seq_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
772 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
773 |
primrec the_only :: "'a seq \<Rightarrow> 'a" where |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
774 |
[code del]: "the_only Empty = undefined" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
775 |
| "the_only (Insert x P) = (if is_empty P then x else let y = singleton P in if x = y then x else undefined)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
776 |
| "the_only (Join P xq) = (if is_empty P then the_only xq else if null xq then singleton P |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
777 |
else let x = singleton P; y = the_only xq in |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
778 |
if x = y then x else undefined)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
779 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
780 |
lemma the_only_singleton: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
781 |
"the_only xq = singleton (pred_of_seq xq)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
782 |
by (induct xq) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
783 |
(auto simp add: singleton_bot singleton_single is_empty_def |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
784 |
null_is_empty Let_def singleton_sup) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
785 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
786 |
lemma singleton_code [code]: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
787 |
"singleton (Seq f) = (case f () |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
788 |
of Empty \<Rightarrow> undefined |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
789 |
| Insert x P \<Rightarrow> if is_empty P then x |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
790 |
else let y = singleton P in |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
791 |
if x = y then x else undefined |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
792 |
| Join P xq \<Rightarrow> if is_empty P then the_only xq |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
793 |
else if null xq then singleton P |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
794 |
else let x = singleton P; y = the_only xq in |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
795 |
if x = y then x else undefined)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
796 |
by (cases "f ()") |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
797 |
(auto simp add: Seq_def the_only_singleton is_empty_def |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
798 |
null_is_empty singleton_bot singleton_single singleton_sup Let_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
799 |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32601
diff
changeset
|
800 |
lemma meta_fun_cong: |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32601
diff
changeset
|
801 |
"f == g ==> f x == g x" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32601
diff
changeset
|
802 |
by simp |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32601
diff
changeset
|
803 |
|
30948 | 804 |
ML {* |
805 |
signature PREDICATE = |
|
806 |
sig |
|
807 |
datatype 'a pred = Seq of (unit -> 'a seq) |
|
808 |
and 'a seq = Empty | Insert of 'a * 'a pred | Join of 'a pred * 'a seq |
|
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
809 |
val yield: 'a pred -> ('a * 'a pred) option |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
810 |
val yieldn: int -> 'a pred -> 'a list * 'a pred |
31222 | 811 |
val map: ('a -> 'b) -> 'a pred -> 'b pred |
30948 | 812 |
end; |
813 |
||
814 |
structure Predicate : PREDICATE = |
|
815 |
struct |
|
816 |
||
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
817 |
@{code_datatype pred = Seq}; |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
818 |
@{code_datatype seq = Empty | Insert | Join}; |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
819 |
|
32372 | 820 |
fun yield (@{code Seq} f) = next (f ()) |
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
821 |
and next @{code Empty} = NONE |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
822 |
| next (@{code Insert} (x, P)) = SOME (x, P) |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
823 |
| next (@{code Join} (P, xq)) = (case yield P |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
824 |
of NONE => next xq |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
825 |
| SOME (x, Q) => SOME (x, @{code Seq} (fn _ => @{code Join} (Q, xq)))) |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
826 |
|
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
827 |
fun anamorph f k x = (if k = 0 then ([], x) |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
828 |
else case f x |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
829 |
of NONE => ([], x) |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
830 |
| SOME (v, y) => let |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
831 |
val (vs, z) = anamorph f (k - 1) y |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
832 |
in (v :: vs, z) end) |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
833 |
|
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
834 |
fun yieldn P = anamorph yield P; |
30948 | 835 |
|
31222 | 836 |
fun map f = @{code map} f; |
837 |
||
30948 | 838 |
end; |
839 |
*} |
|
840 |
||
841 |
code_reserved Eval Predicate |
|
842 |
||
843 |
code_type pred and seq |
|
844 |
(Eval "_/ Predicate.pred" and "_/ Predicate.seq") |
|
845 |
||
846 |
code_const Seq and Empty and Insert and Join |
|
847 |
(Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)") |
|
848 |
||
31122 | 849 |
text {* dummy setup for @{text code_pred} and @{text values} keywords *} |
31108 | 850 |
|
851 |
ML {* |
|
31122 | 852 |
local |
853 |
||
854 |
structure P = OuterParse; |
|
855 |
||
856 |
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; |
|
857 |
||
858 |
in |
|
859 |
||
860 |
val _ = OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate" |
|
861 |
OuterKeyword.thy_goal (P.term_group >> (K (Proof.theorem_i NONE (K I) [[]]))); |
|
862 |
||
31216 | 863 |
val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" |
31122 | 864 |
OuterKeyword.diag ((opt_modes -- P.term) |
865 |
>> (fn (modes, t) => Toplevel.no_timing o Toplevel.keep |
|
866 |
(K ()))); |
|
867 |
||
868 |
end |
|
31108 | 869 |
*} |
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
870 |
|
30328 | 871 |
no_notation |
872 |
inf (infixl "\<sqinter>" 70) and |
|
873 |
sup (infixl "\<squnion>" 65) and |
|
874 |
Inf ("\<Sqinter>_" [900] 900) and |
|
875 |
Sup ("\<Squnion>_" [900] 900) and |
|
876 |
top ("\<top>") and |
|
877 |
bot ("\<bottom>") and |
|
878 |
bind (infixl "\<guillemotright>=" 70) |
|
879 |
||
880 |
hide (open) type pred seq |
|
32582 | 881 |
hide (open) const Pred eval single bind is_empty singleton if_pred not_pred |
882 |
Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map |
|
30328 | 883 |
|
884 |
end |