author | berghofe |
Thu, 12 Jan 2012 00:13:37 +0100 | |
changeset 46191 | a88546428c2a |
parent 46175 | 48c534b22040 |
child 46553 | 50a7e97fe653 |
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 |
41082 | 12 |
bot ("\<bottom>") and |
13 |
top ("\<top>") and |
|
30328 | 14 |
inf (infixl "\<sqinter>" 70) and |
15 |
sup (infixl "\<squnion>" 65) and |
|
16 |
Inf ("\<Sqinter>_" [900] 900) and |
|
41082 | 17 |
Sup ("\<Squnion>_" [900] 900) |
30328 | 18 |
|
41080 | 19 |
syntax (xsymbols) |
41082 | 20 |
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10) |
21 |
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10) |
|
41080 | 22 |
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10) |
23 |
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10) |
|
24 |
||
30328 | 25 |
|
26 |
subsection {* Predicates as (complete) lattices *} |
|
27 |
||
34065
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
28 |
text {* |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
29 |
Handy introduction and elimination rules for @{text "\<le>"} |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
30 |
on unary and binary predicates |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
31 |
*} |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
32 |
|
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
33 |
lemma predicate1I: |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
34 |
assumes PQ: "\<And>x. P x \<Longrightarrow> Q x" |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
35 |
shows "P \<le> Q" |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
36 |
apply (rule le_funI) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
37 |
apply (rule le_boolI) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
38 |
apply (rule PQ) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
39 |
apply assumption |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
40 |
done |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
41 |
|
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
42 |
lemma predicate1D [Pure.dest?, dest?]: |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
43 |
"P \<le> Q \<Longrightarrow> P x \<Longrightarrow> Q x" |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
44 |
apply (erule le_funE) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
45 |
apply (erule le_boolE) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
46 |
apply assumption+ |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
47 |
done |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
48 |
|
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
49 |
lemma rev_predicate1D: |
44414 | 50 |
"P x \<Longrightarrow> P \<le> Q \<Longrightarrow> Q x" |
34065
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
51 |
by (rule predicate1D) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
52 |
|
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
53 |
lemma predicate2I [Pure.intro!, intro!]: |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
54 |
assumes PQ: "\<And>x y. P x y \<Longrightarrow> Q x y" |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
55 |
shows "P \<le> Q" |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
56 |
apply (rule le_funI)+ |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
57 |
apply (rule le_boolI) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
58 |
apply (rule PQ) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
59 |
apply assumption |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
60 |
done |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
61 |
|
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
62 |
lemma predicate2D [Pure.dest, dest]: |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
63 |
"P \<le> Q \<Longrightarrow> P x y \<Longrightarrow> Q x y" |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
64 |
apply (erule le_funE)+ |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
65 |
apply (erule le_boolE) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
66 |
apply assumption+ |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
67 |
done |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
68 |
|
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
69 |
lemma rev_predicate2D: |
44414 | 70 |
"P x y \<Longrightarrow> P \<le> Q \<Longrightarrow> Q x y" |
34065
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
71 |
by (rule predicate2D) |
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
72 |
|
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents:
34007
diff
changeset
|
73 |
|
32779 | 74 |
subsubsection {* Equality *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
75 |
|
46175
48c534b22040
Declared pred_equals/subset_eq, sup_Un_eq and SUP_UN_eq(2) as pred_set_conv rules
berghofe
parents:
46038
diff
changeset
|
76 |
lemma pred_equals_eq [pred_set_conv]: "((\<lambda>x. x \<in> R) = (\<lambda>x. x \<in> S)) \<longleftrightarrow> (R = S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
77 |
by (simp add: set_eq_iff fun_eq_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
78 |
|
44415 | 79 |
lemma pred_equals_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) = (\<lambda>x y. (x, y) \<in> S)) \<longleftrightarrow> (R = S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
80 |
by (simp add: set_eq_iff fun_eq_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
81 |
|
32779 | 82 |
|
83 |
subsubsection {* Order relation *} |
|
84 |
||
46175
48c534b22040
Declared pred_equals/subset_eq, sup_Un_eq and SUP_UN_eq(2) as pred_set_conv rules
berghofe
parents:
46038
diff
changeset
|
85 |
lemma pred_subset_eq [pred_set_conv]: "((\<lambda>x. x \<in> R) \<le> (\<lambda>x. x \<in> S)) \<longleftrightarrow> (R \<subseteq> S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
86 |
by (simp add: subset_iff le_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
87 |
|
44415 | 88 |
lemma pred_subset_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) \<le> (\<lambda>x y. (x, y) \<in> S)) \<longleftrightarrow> (R \<subseteq> S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
89 |
by (simp add: subset_iff le_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
90 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
91 |
|
30328 | 92 |
subsubsection {* Top and bottom elements *} |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
93 |
|
44414 | 94 |
lemma bot1E [no_atp, elim!]: "\<bottom> x \<Longrightarrow> P" |
41550 | 95 |
by (simp add: bot_fun_def) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
96 |
|
44414 | 97 |
lemma bot2E [elim!]: "\<bottom> x y \<Longrightarrow> P" |
41550 | 98 |
by (simp add: bot_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
99 |
|
44414 | 100 |
lemma bot_empty_eq: "\<bottom> = (\<lambda>x. x \<in> {})" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
101 |
by (auto simp add: fun_eq_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
102 |
|
44414 | 103 |
lemma bot_empty_eq2: "\<bottom> = (\<lambda>x y. (x, y) \<in> {})" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
104 |
by (auto simp add: fun_eq_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
105 |
|
44414 | 106 |
lemma top1I [intro!]: "\<top> x" |
41550 | 107 |
by (simp add: top_fun_def) |
41082 | 108 |
|
44414 | 109 |
lemma top2I [intro!]: "\<top> x y" |
41550 | 110 |
by (simp add: top_fun_def) |
41082 | 111 |
|
112 |
||
113 |
subsubsection {* Binary intersection *} |
|
114 |
||
44414 | 115 |
lemma inf1I [intro!]: "A x \<Longrightarrow> B x \<Longrightarrow> (A \<sqinter> B) x" |
41550 | 116 |
by (simp add: inf_fun_def) |
41082 | 117 |
|
44414 | 118 |
lemma inf2I [intro!]: "A x y \<Longrightarrow> B x y \<Longrightarrow> (A \<sqinter> B) x y" |
41550 | 119 |
by (simp add: inf_fun_def) |
41082 | 120 |
|
44414 | 121 |
lemma inf1E [elim!]: "(A \<sqinter> B) x \<Longrightarrow> (A x \<Longrightarrow> B x \<Longrightarrow> P) \<Longrightarrow> P" |
41550 | 122 |
by (simp add: inf_fun_def) |
41082 | 123 |
|
44414 | 124 |
lemma inf2E [elim!]: "(A \<sqinter> B) x y \<Longrightarrow> (A x y \<Longrightarrow> B x y \<Longrightarrow> P) \<Longrightarrow> P" |
41550 | 125 |
by (simp add: inf_fun_def) |
41082 | 126 |
|
44414 | 127 |
lemma inf1D1: "(A \<sqinter> B) x \<Longrightarrow> A x" |
41550 | 128 |
by (simp add: inf_fun_def) |
41082 | 129 |
|
44414 | 130 |
lemma inf2D1: "(A \<sqinter> B) x y \<Longrightarrow> A x y" |
41550 | 131 |
by (simp add: inf_fun_def) |
41082 | 132 |
|
44414 | 133 |
lemma inf1D2: "(A \<sqinter> B) x \<Longrightarrow> B x" |
41550 | 134 |
by (simp add: inf_fun_def) |
41082 | 135 |
|
44414 | 136 |
lemma inf2D2: "(A \<sqinter> B) x y \<Longrightarrow> B x y" |
41550 | 137 |
by (simp add: inf_fun_def) |
41082 | 138 |
|
46191
a88546428c2a
Added inf_Int_eq to pred_set_conv database as well
berghofe
parents:
46175
diff
changeset
|
139 |
lemma inf_Int_eq [pred_set_conv]: "(\<lambda>x. x \<in> R) \<sqinter> (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<inter> S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
140 |
by (simp add: inf_fun_def) |
41082 | 141 |
|
44414 | 142 |
lemma inf_Int_eq2 [pred_set_conv]: "(\<lambda>x y. (x, y) \<in> R) \<sqinter> (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<inter> S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
143 |
by (simp add: inf_fun_def) |
41082 | 144 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
145 |
|
30328 | 146 |
subsubsection {* Binary union *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
147 |
|
45630
0dd654a01217
proper intro? declarations -- NB: elim? indexes on major premise, which is too flexible here and easily leads to inefficiences (cf. a0336f8b6558);
wenzelm
parents:
44928
diff
changeset
|
148 |
lemma sup1I1 [intro?]: "A x \<Longrightarrow> (A \<squnion> B) x" |
41550 | 149 |
by (simp add: sup_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
150 |
|
45630
0dd654a01217
proper intro? declarations -- NB: elim? indexes on major premise, which is too flexible here and easily leads to inefficiences (cf. a0336f8b6558);
wenzelm
parents:
44928
diff
changeset
|
151 |
lemma sup2I1 [intro?]: "A x y \<Longrightarrow> (A \<squnion> B) x y" |
41550 | 152 |
by (simp add: sup_fun_def) |
32883
7cbd93dacef3
inf/sup1/2_iff are mere duplicates of underlying definitions: dropped
haftmann
parents:
32782
diff
changeset
|
153 |
|
45630
0dd654a01217
proper intro? declarations -- NB: elim? indexes on major premise, which is too flexible here and easily leads to inefficiences (cf. a0336f8b6558);
wenzelm
parents:
44928
diff
changeset
|
154 |
lemma sup1I2 [intro?]: "B x \<Longrightarrow> (A \<squnion> B) x" |
41550 | 155 |
by (simp add: sup_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
156 |
|
45630
0dd654a01217
proper intro? declarations -- NB: elim? indexes on major premise, which is too flexible here and easily leads to inefficiences (cf. a0336f8b6558);
wenzelm
parents:
44928
diff
changeset
|
157 |
lemma sup2I2 [intro?]: "B x y \<Longrightarrow> (A \<squnion> B) x y" |
41550 | 158 |
by (simp add: sup_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
159 |
|
44414 | 160 |
lemma sup1E [elim!]: "(A \<squnion> B) x \<Longrightarrow> (A x \<Longrightarrow> P) \<Longrightarrow> (B x \<Longrightarrow> P) \<Longrightarrow> P" |
41550 | 161 |
by (simp add: sup_fun_def) iprover |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
162 |
|
44414 | 163 |
lemma sup2E [elim!]: "(A \<squnion> B) x y \<Longrightarrow> (A x y \<Longrightarrow> P) \<Longrightarrow> (B x y \<Longrightarrow> P) \<Longrightarrow> P" |
41550 | 164 |
by (simp add: sup_fun_def) iprover |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
165 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
166 |
text {* |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
167 |
\medskip Classical introduction rule: no commitment to @{text A} vs |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
168 |
@{text B}. |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
169 |
*} |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
170 |
|
44414 | 171 |
lemma sup1CI [intro!]: "(\<not> B x \<Longrightarrow> A x) \<Longrightarrow> (A \<squnion> B) x" |
41550 | 172 |
by (auto simp add: sup_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
173 |
|
44414 | 174 |
lemma sup2CI [intro!]: "(\<not> B x y \<Longrightarrow> A x y) \<Longrightarrow> (A \<squnion> B) x y" |
41550 | 175 |
by (auto simp add: sup_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
176 |
|
46175
48c534b22040
Declared pred_equals/subset_eq, sup_Un_eq and SUP_UN_eq(2) as pred_set_conv rules
berghofe
parents:
46038
diff
changeset
|
177 |
lemma sup_Un_eq [pred_set_conv]: "(\<lambda>x. x \<in> R) \<squnion> (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<union> S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
178 |
by (simp add: sup_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
179 |
|
44414 | 180 |
lemma sup_Un_eq2 [pred_set_conv]: "(\<lambda>x y. (x, y) \<in> R) \<squnion> (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<union> S)" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
181 |
by (simp add: sup_fun_def) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
182 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
183 |
|
30328 | 184 |
subsubsection {* Intersections of families *} |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
185 |
|
44414 | 186 |
lemma INF1_iff: "(\<Sqinter>x\<in>A. B x) b = (\<forall>x\<in>A. B x b)" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
187 |
by (simp add: INF_apply) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
188 |
|
44414 | 189 |
lemma INF2_iff: "(\<Sqinter>x\<in>A. B x) b c = (\<forall>x\<in>A. B x b c)" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
190 |
by (simp add: INF_apply) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
191 |
|
44414 | 192 |
lemma INF1_I [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> B x b) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
193 |
by (auto simp add: INF_apply) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
194 |
|
44414 | 195 |
lemma INF2_I [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> B x b c) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b c" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
196 |
by (auto simp add: INF_apply) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
197 |
|
44414 | 198 |
lemma INF1_D [elim]: "(\<Sqinter>x\<in>A. B x) b \<Longrightarrow> a \<in> A \<Longrightarrow> B a b" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
199 |
by (auto simp add: INF_apply) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
200 |
|
44414 | 201 |
lemma INF2_D [elim]: "(\<Sqinter>x\<in>A. B x) b c \<Longrightarrow> a \<in> A \<Longrightarrow> B a b c" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
202 |
by (auto simp add: INF_apply) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
203 |
|
44414 | 204 |
lemma INF1_E [elim]: "(\<Sqinter>x\<in>A. B x) b \<Longrightarrow> (B a b \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
205 |
by (auto simp add: INF_apply) |
22430
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
berghofe
parents:
22422
diff
changeset
|
206 |
|
44414 | 207 |
lemma INF2_E [elim]: "(\<Sqinter>x\<in>A. B x) b c \<Longrightarrow> (B a b c \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
208 |
by (auto simp add: INF_apply) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
209 |
|
44414 | 210 |
lemma INF_INT_eq: "(\<Sqinter>i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (\<Sqinter>i. r i))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
211 |
by (simp add: INF_apply fun_eq_iff) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
212 |
|
44414 | 213 |
lemma INF_INT_eq2: "(\<Sqinter>i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (\<Sqinter>i. r i))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
214 |
by (simp add: INF_apply fun_eq_iff) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
215 |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
216 |
|
41082 | 217 |
subsubsection {* Unions of families *} |
218 |
||
44414 | 219 |
lemma SUP1_iff: "(\<Squnion>x\<in>A. B x) b = (\<exists>x\<in>A. B x b)" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
220 |
by (simp add: SUP_apply) |
41082 | 221 |
|
44414 | 222 |
lemma SUP2_iff: "(\<Squnion>x\<in>A. B x) b c = (\<exists>x\<in>A. B x b c)" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
223 |
by (simp add: SUP_apply) |
41082 | 224 |
|
44414 | 225 |
lemma SUP1_I [intro]: "a \<in> A \<Longrightarrow> B a b \<Longrightarrow> (\<Squnion>x\<in>A. B x) b" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
226 |
by (auto simp add: SUP_apply) |
41082 | 227 |
|
44414 | 228 |
lemma SUP2_I [intro]: "a \<in> A \<Longrightarrow> B a b c \<Longrightarrow> (\<Squnion>x\<in>A. B x) b c" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
229 |
by (auto simp add: SUP_apply) |
41082 | 230 |
|
44414 | 231 |
lemma SUP1_E [elim!]: "(\<Squnion>x\<in>A. B x) b \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> B x b \<Longrightarrow> R) \<Longrightarrow> R" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
232 |
by (auto simp add: SUP_apply) |
41082 | 233 |
|
44414 | 234 |
lemma SUP2_E [elim!]: "(\<Squnion>x\<in>A. B x) b c \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> B x b c \<Longrightarrow> R) \<Longrightarrow> R" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
235 |
by (auto simp add: SUP_apply) |
41082 | 236 |
|
46175
48c534b22040
Declared pred_equals/subset_eq, sup_Un_eq and SUP_UN_eq(2) as pred_set_conv rules
berghofe
parents:
46038
diff
changeset
|
237 |
lemma SUP_UN_eq [pred_set_conv]: "(\<Squnion>i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (\<Union>i. r i))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
238 |
by (simp add: SUP_apply fun_eq_iff) |
41082 | 239 |
|
46175
48c534b22040
Declared pred_equals/subset_eq, sup_Un_eq and SUP_UN_eq(2) as pred_set_conv rules
berghofe
parents:
46038
diff
changeset
|
240 |
lemma SUP_UN_eq2 [pred_set_conv]: "(\<Squnion>i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (\<Union>i. r i))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
241 |
by (simp add: SUP_apply fun_eq_iff) |
41082 | 242 |
|
243 |
||
30328 | 244 |
subsection {* Predicates as relations *} |
245 |
||
246 |
subsubsection {* Composition *} |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
247 |
|
44414 | 248 |
inductive pred_comp :: "['a \<Rightarrow> 'b \<Rightarrow> bool, 'b \<Rightarrow> 'c \<Rightarrow> bool] \<Rightarrow> 'a \<Rightarrow> 'c \<Rightarrow> bool" (infixr "OO" 75) |
249 |
for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and s :: "'b \<Rightarrow> 'c \<Rightarrow> bool" where |
|
250 |
pred_compI [intro]: "r a b \<Longrightarrow> s b c \<Longrightarrow> (r OO s) a c" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
251 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
252 |
inductive_cases pred_compE [elim!]: "(r OO s) a c" |
22259
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 pred_comp_rel_comp_eq [pred_set_conv]: |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
255 |
"((\<lambda>x y. (x, y) \<in> r) OO (\<lambda>x y. (x, y) \<in> s)) = (\<lambda>x y. (x, y) \<in> r O s)" |
41550 | 256 |
by (auto simp add: fun_eq_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
257 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
258 |
|
30328 | 259 |
subsubsection {* Converse *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
260 |
|
44414 | 261 |
inductive conversep :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> bool" ("(_^--1)" [1000] 1000) |
262 |
for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where |
|
263 |
conversepI: "r a b \<Longrightarrow> r^--1 b a" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
264 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
265 |
notation (xsymbols) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
266 |
conversep ("(_\<inverse>\<inverse>)" [1000] 1000) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
267 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
268 |
lemma conversepD: |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
269 |
assumes ab: "r^--1 a b" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
270 |
shows "r b a" using ab |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
271 |
by cases simp |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
272 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
273 |
lemma conversep_iff [iff]: "r^--1 a b = r b a" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
274 |
by (iprover intro: conversepI 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_converse_eq [pred_set_conv]: |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
277 |
"(\<lambda>x y. (x, y) \<in> r)^--1 = (\<lambda>x y. (x, y) \<in> r^-1)" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
278 |
by (auto simp add: fun_eq_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
279 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
280 |
lemma conversep_conversep [simp]: "(r^--1)^--1 = r" |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
281 |
by (iprover intro: order_antisym conversepI dest: conversepD) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
282 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
283 |
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
|
284 |
by (iprover intro: order_antisym conversepI pred_compI |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
285 |
elim: pred_compE dest: conversepD) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
286 |
|
44414 | 287 |
lemma converse_meet: "(r \<sqinter> s)^--1 = r^--1 \<sqinter> s^--1" |
41550 | 288 |
by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
289 |
|
44414 | 290 |
lemma converse_join: "(r \<squnion> s)^--1 = r^--1 \<squnion> s^--1" |
41550 | 291 |
by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
292 |
|
44414 | 293 |
lemma conversep_noteq [simp]: "(op \<noteq>)^--1 = op \<noteq>" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
294 |
by (auto simp add: fun_eq_iff) |
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 |
lemma conversep_eq [simp]: "(op =)^--1 = op =" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
297 |
by (auto simp add: fun_eq_iff) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
298 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
299 |
|
30328 | 300 |
subsubsection {* Domain *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
301 |
|
44414 | 302 |
inductive DomainP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" |
303 |
for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where |
|
304 |
DomainPI [intro]: "r a b \<Longrightarrow> DomainP r a" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
305 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
306 |
inductive_cases DomainPE [elim!]: "DomainP r a" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
307 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
308 |
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
|
309 |
by (blast intro!: Orderings.order_antisym predicate1I) |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
310 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
311 |
|
30328 | 312 |
subsubsection {* Range *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
313 |
|
44414 | 314 |
inductive RangeP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool" |
315 |
for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where |
|
316 |
RangePI [intro]: "r a b \<Longrightarrow> RangeP r b" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
317 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
318 |
inductive_cases RangePE [elim!]: "RangeP r b" |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
319 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
320 |
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
|
321 |
by (blast intro!: Orderings.order_antisym predicate1I) |
22259
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 {* Inverse image *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
325 |
|
44414 | 326 |
definition inv_imagep :: "('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where |
327 |
"inv_imagep r f = (\<lambda>x y. r (f x) (f y))" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
328 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
329 |
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
|
330 |
by (simp add: inv_image_def inv_imagep_def) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
331 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
332 |
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
|
333 |
by (simp add: inv_imagep_def) |
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
334 |
|
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
335 |
|
30328 | 336 |
subsubsection {* Powerset *} |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
337 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
338 |
definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where |
44414 | 339 |
"Powp A = (\<lambda>B. \<forall>x \<in> B. A x)" |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
340 |
|
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
341 |
lemma Powp_Pow_eq [pred_set_conv]: "Powp (\<lambda>x. x \<in> A) = (\<lambda>x. x \<in> Pow A)" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
342 |
by (auto simp add: Powp_def fun_eq_iff) |
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
343 |
|
46175
48c534b22040
Declared pred_equals/subset_eq, sup_Un_eq and SUP_UN_eq(2) as pred_set_conv rules
berghofe
parents:
46038
diff
changeset
|
344 |
lemmas Powp_mono [mono] = Pow_mono [to_pred] |
26797
a6cb51c314f2
- Added mem_def and predicate1I in some of the proofs
berghofe
parents:
24345
diff
changeset
|
345 |
|
23741
1801a921df13
- Moved infrastructure for converting between sets and predicates
berghofe
parents:
23708
diff
changeset
|
346 |
|
30328 | 347 |
subsubsection {* Properties of relations *} |
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
348 |
|
44414 | 349 |
abbreviation antisymP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where |
350 |
"antisymP r \<equiv> antisym {(x, y). r x y}" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
351 |
|
44414 | 352 |
abbreviation transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where |
353 |
"transP r \<equiv> trans {(x, y). r x y}" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
354 |
|
44414 | 355 |
abbreviation single_valuedP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where |
356 |
"single_valuedP r \<equiv> single_valued {(x, y). r x y}" |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
357 |
|
40813
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
358 |
(*FIXME inconsistencies: abbreviations vs. definitions, suffix `P` vs. suffix `p`*) |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
359 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
360 |
definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
361 |
"reflp r \<longleftrightarrow> refl {(x, y). r x y}" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
362 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
363 |
definition symp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
364 |
"symp r \<longleftrightarrow> sym {(x, y). r x y}" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
365 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
366 |
definition transp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
367 |
"transp r \<longleftrightarrow> trans {(x, y). r x y}" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
368 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
369 |
lemma reflpI: |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
370 |
"(\<And>x. r x x) \<Longrightarrow> reflp r" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
371 |
by (auto intro: refl_onI simp add: reflp_def) |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
372 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
373 |
lemma reflpE: |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
374 |
assumes "reflp r" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
375 |
obtains "r x x" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
376 |
using assms by (auto dest: refl_onD simp add: reflp_def) |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
377 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
378 |
lemma sympI: |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
379 |
"(\<And>x y. r x y \<Longrightarrow> r y x) \<Longrightarrow> symp r" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
380 |
by (auto intro: symI simp add: symp_def) |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
381 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
382 |
lemma sympE: |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
383 |
assumes "symp r" and "r x y" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
384 |
obtains "r y x" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
385 |
using assms by (auto dest: symD simp add: symp_def) |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
386 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
387 |
lemma transpI: |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
388 |
"(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
389 |
by (auto intro: transI simp add: transp_def) |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
390 |
|
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
391 |
lemma transpE: |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
392 |
assumes "transp r" and "r x y" and "r y z" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
393 |
obtains "r x z" |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
394 |
using assms by (auto dest: transD simp add: transp_def) |
f1fc2a1547eb
moved generic definitions about relations from Quotient.thy to Predicate;
haftmann
parents:
40674
diff
changeset
|
395 |
|
30328 | 396 |
|
397 |
subsection {* Predicates as enumerations *} |
|
398 |
||
399 |
subsubsection {* The type of predicate enumerations (a monad) *} |
|
400 |
||
401 |
datatype 'a pred = Pred "'a \<Rightarrow> bool" |
|
402 |
||
403 |
primrec eval :: "'a pred \<Rightarrow> 'a \<Rightarrow> bool" where |
|
404 |
eval_pred: "eval (Pred f) = f" |
|
405 |
||
406 |
lemma Pred_eval [simp]: |
|
407 |
"Pred (eval x) = x" |
|
408 |
by (cases x) simp |
|
409 |
||
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
410 |
lemma pred_eqI: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
411 |
"(\<And>w. eval P w \<longleftrightarrow> eval Q w) \<Longrightarrow> P = Q" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
412 |
by (cases P, cases Q) (auto simp add: fun_eq_iff) |
30328 | 413 |
|
46038
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
414 |
lemma pred_eq_iff: |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
415 |
"P = Q \<Longrightarrow> (\<And>w. eval P w \<longleftrightarrow> eval Q w)" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
416 |
by (simp add: pred_eqI) |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
417 |
|
44033 | 418 |
instantiation pred :: (type) complete_lattice |
30328 | 419 |
begin |
420 |
||
421 |
definition |
|
422 |
"P \<le> Q \<longleftrightarrow> eval P \<le> eval Q" |
|
423 |
||
424 |
definition |
|
425 |
"P < Q \<longleftrightarrow> eval P < eval Q" |
|
426 |
||
427 |
definition |
|
428 |
"\<bottom> = Pred \<bottom>" |
|
429 |
||
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
430 |
lemma eval_bot [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
431 |
"eval \<bottom> = \<bottom>" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
432 |
by (simp add: bot_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
433 |
|
30328 | 434 |
definition |
435 |
"\<top> = Pred \<top>" |
|
436 |
||
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
437 |
lemma eval_top [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
438 |
"eval \<top> = \<top>" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
439 |
by (simp add: top_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
440 |
|
30328 | 441 |
definition |
442 |
"P \<sqinter> Q = Pred (eval P \<sqinter> eval Q)" |
|
443 |
||
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
444 |
lemma eval_inf [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
445 |
"eval (P \<sqinter> Q) = eval P \<sqinter> eval Q" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
446 |
by (simp add: inf_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
447 |
|
30328 | 448 |
definition |
449 |
"P \<squnion> Q = Pred (eval P \<squnion> eval Q)" |
|
450 |
||
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
451 |
lemma eval_sup [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
452 |
"eval (P \<squnion> Q) = eval P \<squnion> eval Q" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
453 |
by (simp add: sup_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
454 |
|
30328 | 455 |
definition |
37767 | 456 |
"\<Sqinter>A = Pred (INFI A eval)" |
30328 | 457 |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
458 |
lemma eval_Inf [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
459 |
"eval (\<Sqinter>A) = INFI A eval" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
460 |
by (simp add: Inf_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
461 |
|
30328 | 462 |
definition |
37767 | 463 |
"\<Squnion>A = Pred (SUPR A eval)" |
30328 | 464 |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
465 |
lemma eval_Sup [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
466 |
"eval (\<Squnion>A) = SUPR A eval" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
467 |
by (simp add: Sup_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
468 |
|
44033 | 469 |
instance proof |
44415 | 470 |
qed (auto intro!: pred_eqI simp add: less_eq_pred_def less_pred_def le_fun_def less_fun_def) |
44033 | 471 |
|
472 |
end |
|
473 |
||
474 |
lemma eval_INFI [simp]: |
|
475 |
"eval (INFI A f) = INFI A (eval \<circ> f)" |
|
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
476 |
by (simp only: INF_def eval_Inf image_compose) |
44033 | 477 |
|
478 |
lemma eval_SUPR [simp]: |
|
479 |
"eval (SUPR A f) = SUPR A (eval \<circ> f)" |
|
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44415
diff
changeset
|
480 |
by (simp only: SUP_def eval_Sup image_compose) |
44033 | 481 |
|
482 |
instantiation pred :: (type) complete_boolean_algebra |
|
483 |
begin |
|
484 |
||
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
485 |
definition |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
486 |
"- P = Pred (- eval P)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
487 |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
488 |
lemma eval_compl [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
489 |
"eval (- P) = - eval P" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
490 |
by (simp add: uminus_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
491 |
|
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
492 |
definition |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
493 |
"P - Q = Pred (eval P - eval Q)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
494 |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
495 |
lemma eval_minus [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
496 |
"eval (P - Q) = eval P - eval Q" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
497 |
by (simp add: minus_pred_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
498 |
|
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
499 |
instance proof |
44415 | 500 |
qed (auto intro!: pred_eqI simp add: uminus_apply minus_apply INF_apply SUP_apply) |
30328 | 501 |
|
22259
476604be7d88
New theory for converting between predicates and sets.
berghofe
parents:
diff
changeset
|
502 |
end |
30328 | 503 |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
504 |
definition single :: "'a \<Rightarrow> 'a pred" where |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
505 |
"single x = Pred ((op =) x)" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
506 |
|
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
507 |
lemma eval_single [simp]: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
508 |
"eval (single x) = (op =) x" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
509 |
by (simp add: single_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
510 |
|
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
511 |
definition bind :: "'a pred \<Rightarrow> ('a \<Rightarrow> 'b pred) \<Rightarrow> 'b pred" (infixl "\<guillemotright>=" 70) where |
41080 | 512 |
"P \<guillemotright>= f = (SUPR {x. eval P x} f)" |
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
513 |
|
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
514 |
lemma eval_bind [simp]: |
41080 | 515 |
"eval (P \<guillemotright>= f) = eval (SUPR {x. eval P x} f)" |
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
516 |
by (simp add: bind_def) |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
517 |
|
30328 | 518 |
lemma bind_bind: |
519 |
"(P \<guillemotright>= Q) \<guillemotright>= R = P \<guillemotright>= (\<lambda>x. Q x \<guillemotright>= R)" |
|
44415 | 520 |
by (rule pred_eqI) (auto simp add: SUP_apply) |
30328 | 521 |
|
522 |
lemma bind_single: |
|
523 |
"P \<guillemotright>= single = P" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
524 |
by (rule pred_eqI) auto |
30328 | 525 |
|
526 |
lemma single_bind: |
|
527 |
"single x \<guillemotright>= P = P x" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
528 |
by (rule pred_eqI) auto |
30328 | 529 |
|
530 |
lemma bottom_bind: |
|
531 |
"\<bottom> \<guillemotright>= P = \<bottom>" |
|
40674
54dbe6a1c349
adhere established Collect/mem convention more closely
haftmann
parents:
40616
diff
changeset
|
532 |
by (rule pred_eqI) auto |
30328 | 533 |
|
534 |
lemma sup_bind: |
|
535 |
"(P \<squnion> Q) \<guillemotright>= R = P \<guillemotright>= R \<squnion> Q \<guillemotright>= R" |
|
40674
54dbe6a1c349
adhere established Collect/mem convention more closely
haftmann
parents:
40616
diff
changeset
|
536 |
by (rule pred_eqI) auto |
30328 | 537 |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
538 |
lemma Sup_bind: |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
539 |
"(\<Squnion>A \<guillemotright>= f) = \<Squnion>((\<lambda>x. x \<guillemotright>= f) ` A)" |
44415 | 540 |
by (rule pred_eqI) (auto simp add: SUP_apply) |
30328 | 541 |
|
542 |
lemma pred_iffI: |
|
543 |
assumes "\<And>x. eval A x \<Longrightarrow> eval B x" |
|
544 |
and "\<And>x. eval B x \<Longrightarrow> eval A x" |
|
545 |
shows "A = B" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
546 |
using assms by (auto intro: pred_eqI) |
30328 | 547 |
|
548 |
lemma singleI: "eval (single x) x" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
549 |
by simp |
30328 | 550 |
|
551 |
lemma singleI_unit: "eval (single ()) x" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
552 |
by simp |
30328 | 553 |
|
554 |
lemma singleE: "eval (single x) y \<Longrightarrow> (y = x \<Longrightarrow> P) \<Longrightarrow> P" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
555 |
by simp |
30328 | 556 |
|
557 |
lemma singleE': "eval (single x) y \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> P" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
558 |
by simp |
30328 | 559 |
|
560 |
lemma bindI: "eval P x \<Longrightarrow> eval (Q x) y \<Longrightarrow> eval (P \<guillemotright>= Q) y" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
561 |
by auto |
30328 | 562 |
|
563 |
lemma bindE: "eval (R \<guillemotright>= Q) y \<Longrightarrow> (\<And>x. eval R x \<Longrightarrow> eval (Q x) y \<Longrightarrow> P) \<Longrightarrow> P" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
564 |
by auto |
30328 | 565 |
|
566 |
lemma botE: "eval \<bottom> x \<Longrightarrow> P" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
567 |
by auto |
30328 | 568 |
|
569 |
lemma supI1: "eval A x \<Longrightarrow> eval (A \<squnion> B) x" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
570 |
by auto |
30328 | 571 |
|
572 |
lemma supI2: "eval B x \<Longrightarrow> eval (A \<squnion> B) x" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
573 |
by auto |
30328 | 574 |
|
575 |
lemma supE: "eval (A \<squnion> B) x \<Longrightarrow> (eval A x \<Longrightarrow> P) \<Longrightarrow> (eval B x \<Longrightarrow> P) \<Longrightarrow> P" |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
576 |
by auto |
30328 | 577 |
|
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
578 |
lemma single_not_bot [simp]: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
579 |
"single x \<noteq> \<bottom>" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
580 |
by (auto simp add: single_def bot_pred_def fun_eq_iff) |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
581 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
582 |
lemma not_bot: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
583 |
assumes "A \<noteq> \<bottom>" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
584 |
obtains x where "eval A x" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
585 |
using assms by (cases A) (auto simp add: bot_pred_def) |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45630
diff
changeset
|
586 |
|
32578
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 |
subsubsection {* Emptiness check and definite choice *} |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
589 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
590 |
definition is_empty :: "'a pred \<Rightarrow> bool" where |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
591 |
"is_empty A \<longleftrightarrow> A = \<bottom>" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
592 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
593 |
lemma is_empty_bot: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
594 |
"is_empty \<bottom>" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
595 |
by (simp add: is_empty_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
596 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
597 |
lemma not_is_empty_single: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
598 |
"\<not> is_empty (single x)" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
599 |
by (auto simp add: is_empty_def single_def bot_pred_def fun_eq_iff) |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
600 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
601 |
lemma is_empty_sup: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
602 |
"is_empty (A \<squnion> B) \<longleftrightarrow> is_empty A \<and> is_empty B" |
36008 | 603 |
by (auto simp add: is_empty_def) |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
604 |
|
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
605 |
definition singleton :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a pred \<Rightarrow> 'a" where |
33111 | 606 |
"singleton dfault A = (if \<exists>!x. eval A x then THE x. eval A x else dfault ())" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
607 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
608 |
lemma singleton_eqI: |
33110 | 609 |
"\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton dfault A = x" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
610 |
by (auto simp add: singleton_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
611 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
612 |
lemma eval_singletonI: |
33110 | 613 |
"\<exists>!x. eval A x \<Longrightarrow> eval A (singleton dfault A)" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
614 |
proof - |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
615 |
assume assm: "\<exists>!x. eval A x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
616 |
then obtain x where "eval A x" .. |
33110 | 617 |
moreover with assm have "singleton dfault A = x" by (rule singleton_eqI) |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
618 |
ultimately show ?thesis by simp |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
619 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
620 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
621 |
lemma single_singleton: |
33110 | 622 |
"\<exists>!x. eval A x \<Longrightarrow> single (singleton dfault A) = A" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
623 |
proof - |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
624 |
assume assm: "\<exists>!x. eval A x" |
33110 | 625 |
then have "eval A (singleton dfault A)" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
626 |
by (rule eval_singletonI) |
33110 | 627 |
moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton dfault A = x" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
628 |
by (rule singleton_eqI) |
33110 | 629 |
ultimately have "eval (single (singleton dfault A)) = eval A" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
630 |
by (simp (no_asm_use) add: single_def fun_eq_iff) blast |
40616
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
631 |
then have "\<And>x. eval (single (singleton dfault A)) x = eval A x" |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
632 |
by simp |
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
haftmann
parents:
39302
diff
changeset
|
633 |
then show ?thesis by (rule pred_eqI) |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
634 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
635 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
636 |
lemma singleton_undefinedI: |
33111 | 637 |
"\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton dfault A = dfault ()" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
638 |
by (simp add: singleton_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
639 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
640 |
lemma singleton_bot: |
33111 | 641 |
"singleton dfault \<bottom> = dfault ()" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
642 |
by (auto simp add: bot_pred_def intro: singleton_undefinedI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
643 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
644 |
lemma singleton_single: |
33110 | 645 |
"singleton dfault (single x) = x" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
646 |
by (auto simp add: intro: singleton_eqI singleI elim: singleE) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
647 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
648 |
lemma singleton_sup_single_single: |
33111 | 649 |
"singleton dfault (single x \<squnion> single y) = (if x = y then x else dfault ())" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
650 |
proof (cases "x = y") |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
651 |
case True then show ?thesis by (simp add: singleton_single) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
652 |
next |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
653 |
case False |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
654 |
have "eval (single x \<squnion> single y) x" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
655 |
and "eval (single x \<squnion> single y) y" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
656 |
by (auto intro: supI1 supI2 singleI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
657 |
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
|
658 |
by blast |
33111 | 659 |
then have "singleton dfault (single x \<squnion> single y) = dfault ()" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
660 |
by (rule singleton_undefinedI) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
661 |
with False show ?thesis by simp |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
662 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
663 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
664 |
lemma singleton_sup_aux: |
33110 | 665 |
"singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B |
666 |
else if B = \<bottom> then singleton dfault A |
|
667 |
else singleton dfault |
|
668 |
(single (singleton dfault A) \<squnion> single (singleton dfault B)))" |
|
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
669 |
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
|
670 |
case True then show ?thesis by (simp add: single_singleton) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
671 |
next |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
672 |
case False |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
673 |
from False have A_or_B: |
33111 | 674 |
"singleton dfault A = dfault () \<or> singleton dfault B = dfault ()" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
675 |
by (auto intro!: singleton_undefinedI) |
33110 | 676 |
then have rhs: "singleton dfault |
33111 | 677 |
(single (singleton dfault A) \<squnion> single (singleton dfault B)) = dfault ()" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
678 |
by (auto simp add: singleton_sup_single_single singleton_single) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
679 |
from False have not_unique: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
680 |
"\<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
|
681 |
show ?thesis proof (cases "A \<noteq> \<bottom> \<and> B \<noteq> \<bottom>") |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
682 |
case True |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
683 |
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
|
684 |
by (blast elim: not_bot) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
685 |
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
|
686 |
by (auto simp add: sup_pred_def bot_pred_def) |
33111 | 687 |
then have "singleton dfault (A \<squnion> B) = dfault ()" by (rule singleton_undefinedI) |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
688 |
with True rhs show ?thesis by simp |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
689 |
next |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
690 |
case False then show ?thesis by auto |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
691 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
692 |
qed |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
693 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
694 |
lemma singleton_sup: |
33110 | 695 |
"singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B |
696 |
else if B = \<bottom> then singleton dfault A |
|
33111 | 697 |
else if singleton dfault A = singleton dfault B then singleton dfault A else dfault ())" |
33110 | 698 |
using singleton_sup_aux [of dfault A B] by (simp only: singleton_sup_single_single) |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
699 |
|
30328 | 700 |
|
701 |
subsubsection {* Derived operations *} |
|
702 |
||
703 |
definition if_pred :: "bool \<Rightarrow> unit pred" where |
|
704 |
if_pred_eq: "if_pred b = (if b then single () else \<bottom>)" |
|
705 |
||
33754
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
706 |
definition holds :: "unit pred \<Rightarrow> bool" where |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
707 |
holds_eq: "holds P = eval P ()" |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
708 |
|
30328 | 709 |
definition not_pred :: "unit pred \<Rightarrow> unit pred" where |
710 |
not_pred_eq: "not_pred P = (if eval P () then \<bottom> else single ())" |
|
711 |
||
712 |
lemma if_predI: "P \<Longrightarrow> eval (if_pred P) ()" |
|
713 |
unfolding if_pred_eq by (auto intro: singleI) |
|
714 |
||
715 |
lemma if_predE: "eval (if_pred b) x \<Longrightarrow> (b \<Longrightarrow> x = () \<Longrightarrow> P) \<Longrightarrow> P" |
|
716 |
unfolding if_pred_eq by (cases b) (auto elim: botE) |
|
717 |
||
718 |
lemma not_predI: "\<not> P \<Longrightarrow> eval (not_pred (Pred (\<lambda>u. P))) ()" |
|
719 |
unfolding not_pred_eq eval_pred by (auto intro: singleI) |
|
720 |
||
721 |
lemma not_predI': "\<not> eval P () \<Longrightarrow> eval (not_pred P) ()" |
|
722 |
unfolding not_pred_eq by (auto intro: singleI) |
|
723 |
||
724 |
lemma not_predE: "eval (not_pred (Pred (\<lambda>u. P))) x \<Longrightarrow> (\<not> P \<Longrightarrow> thesis) \<Longrightarrow> thesis" |
|
725 |
unfolding not_pred_eq |
|
726 |
by (auto split: split_if_asm elim: botE) |
|
727 |
||
728 |
lemma not_predE': "eval (not_pred P) x \<Longrightarrow> (\<not> eval P x \<Longrightarrow> thesis) \<Longrightarrow> thesis" |
|
729 |
unfolding not_pred_eq |
|
730 |
by (auto split: split_if_asm elim: botE) |
|
33754
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
731 |
lemma "f () = False \<or> f () = True" |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
732 |
by simp |
30328 | 733 |
|
37549 | 734 |
lemma closure_of_bool_cases [no_atp]: |
44007 | 735 |
fixes f :: "unit \<Rightarrow> bool" |
736 |
assumes "f = (\<lambda>u. False) \<Longrightarrow> P f" |
|
737 |
assumes "f = (\<lambda>u. True) \<Longrightarrow> P f" |
|
738 |
shows "P f" |
|
33754
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
739 |
proof - |
44007 | 740 |
have "f = (\<lambda>u. False) \<or> f = (\<lambda>u. True)" |
33754
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
741 |
apply (cases "f ()") |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
742 |
apply (rule disjI2) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
743 |
apply (rule ext) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
744 |
apply (simp add: unit_eq) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
745 |
apply (rule disjI1) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
746 |
apply (rule ext) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
747 |
apply (simp add: unit_eq) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
748 |
done |
41550 | 749 |
from this assms show ?thesis by blast |
33754
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
750 |
qed |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
751 |
|
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
752 |
lemma unit_pred_cases: |
44007 | 753 |
assumes "P \<bottom>" |
754 |
assumes "P (single ())" |
|
755 |
shows "P Q" |
|
44415 | 756 |
using assms unfolding bot_pred_def bot_fun_def bot_bool_def empty_def single_def proof (cases Q) |
44007 | 757 |
fix f |
758 |
assume "P (Pred (\<lambda>u. False))" "P (Pred (\<lambda>u. () = u))" |
|
759 |
then have "P (Pred f)" |
|
760 |
by (cases _ f rule: closure_of_bool_cases) simp_all |
|
761 |
moreover assume "Q = Pred f" |
|
762 |
ultimately show "P Q" by simp |
|
763 |
qed |
|
764 |
||
33754
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
765 |
lemma holds_if_pred: |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
766 |
"holds (if_pred b) = b" |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
767 |
unfolding if_pred_eq holds_eq |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
768 |
by (cases b) (auto intro: singleI elim: botE) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
769 |
|
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
770 |
lemma if_pred_holds: |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
771 |
"if_pred (holds P) = P" |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
772 |
unfolding if_pred_eq holds_eq |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
773 |
by (rule unit_pred_cases) (auto intro: singleI elim: botE) |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
774 |
|
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
775 |
lemma is_empty_holds: |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
776 |
"is_empty P \<longleftrightarrow> \<not> holds P" |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
777 |
unfolding is_empty_def holds_eq |
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33622
diff
changeset
|
778 |
by (rule unit_pred_cases) (auto elim: botE intro: singleI) |
30328 | 779 |
|
41311 | 780 |
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pred \<Rightarrow> 'b pred" where |
781 |
"map f P = P \<guillemotright>= (single o f)" |
|
782 |
||
783 |
lemma eval_map [simp]: |
|
44363 | 784 |
"eval (map f P) = (\<Squnion>x\<in>{x. eval P x}. (\<lambda>y. f x = y))" |
44415 | 785 |
by (auto simp add: map_def comp_def) |
41311 | 786 |
|
41505
6d19301074cf
"enriched_type" replaces less specific "type_lifting"
haftmann
parents:
41372
diff
changeset
|
787 |
enriched_type map: map |
44363 | 788 |
by (rule ext, rule pred_eqI, auto)+ |
41311 | 789 |
|
790 |
||
30328 | 791 |
subsubsection {* Implementation *} |
792 |
||
793 |
datatype 'a seq = Empty | Insert "'a" "'a pred" | Join "'a pred" "'a seq" |
|
794 |
||
795 |
primrec pred_of_seq :: "'a seq \<Rightarrow> 'a pred" where |
|
44414 | 796 |
"pred_of_seq Empty = \<bottom>" |
797 |
| "pred_of_seq (Insert x P) = single x \<squnion> P" |
|
798 |
| "pred_of_seq (Join P xq) = P \<squnion> pred_of_seq xq" |
|
30328 | 799 |
|
800 |
definition Seq :: "(unit \<Rightarrow> 'a seq) \<Rightarrow> 'a pred" where |
|
801 |
"Seq f = pred_of_seq (f ())" |
|
802 |
||
803 |
code_datatype Seq |
|
804 |
||
805 |
primrec member :: "'a seq \<Rightarrow> 'a \<Rightarrow> bool" where |
|
806 |
"member Empty x \<longleftrightarrow> False" |
|
44414 | 807 |
| "member (Insert y P) x \<longleftrightarrow> x = y \<or> eval P x" |
808 |
| "member (Join P xq) x \<longleftrightarrow> eval P x \<or> member xq x" |
|
30328 | 809 |
|
810 |
lemma eval_member: |
|
811 |
"member xq = eval (pred_of_seq xq)" |
|
812 |
proof (induct xq) |
|
813 |
case Empty show ?case |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
814 |
by (auto simp add: fun_eq_iff elim: botE) |
30328 | 815 |
next |
816 |
case Insert show ?case |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
817 |
by (auto simp add: fun_eq_iff elim: supE singleE intro: supI1 supI2 singleI) |
30328 | 818 |
next |
819 |
case Join then show ?case |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
820 |
by (auto simp add: fun_eq_iff elim: supE intro: supI1 supI2) |
30328 | 821 |
qed |
822 |
||
46038
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
823 |
lemma eval_code [(* FIXME declare simp *)code]: "eval (Seq f) = member (f ())" |
30328 | 824 |
unfolding Seq_def by (rule sym, rule eval_member) |
825 |
||
826 |
lemma single_code [code]: |
|
827 |
"single x = Seq (\<lambda>u. Insert x \<bottom>)" |
|
828 |
unfolding Seq_def by simp |
|
829 |
||
41080 | 830 |
primrec "apply" :: "('a \<Rightarrow> 'b pred) \<Rightarrow> 'a seq \<Rightarrow> 'b seq" where |
44415 | 831 |
"apply f Empty = Empty" |
832 |
| "apply f (Insert x P) = Join (f x) (Join (P \<guillemotright>= f) Empty)" |
|
833 |
| "apply f (Join P xq) = Join (P \<guillemotright>= f) (apply f xq)" |
|
30328 | 834 |
|
835 |
lemma apply_bind: |
|
836 |
"pred_of_seq (apply f xq) = pred_of_seq xq \<guillemotright>= f" |
|
837 |
proof (induct xq) |
|
838 |
case Empty show ?case |
|
839 |
by (simp add: bottom_bind) |
|
840 |
next |
|
841 |
case Insert show ?case |
|
842 |
by (simp add: single_bind sup_bind) |
|
843 |
next |
|
844 |
case Join then show ?case |
|
845 |
by (simp add: sup_bind) |
|
846 |
qed |
|
847 |
||
848 |
lemma bind_code [code]: |
|
849 |
"Seq g \<guillemotright>= f = Seq (\<lambda>u. apply f (g ()))" |
|
850 |
unfolding Seq_def by (rule sym, rule apply_bind) |
|
851 |
||
852 |
lemma bot_set_code [code]: |
|
853 |
"\<bottom> = Seq (\<lambda>u. Empty)" |
|
854 |
unfolding Seq_def by simp |
|
855 |
||
30376 | 856 |
primrec adjunct :: "'a pred \<Rightarrow> 'a seq \<Rightarrow> 'a seq" where |
44415 | 857 |
"adjunct P Empty = Join P Empty" |
858 |
| "adjunct P (Insert x Q) = Insert x (Q \<squnion> P)" |
|
859 |
| "adjunct P (Join Q xq) = Join Q (adjunct P xq)" |
|
30376 | 860 |
|
861 |
lemma adjunct_sup: |
|
862 |
"pred_of_seq (adjunct P xq) = P \<squnion> pred_of_seq xq" |
|
863 |
by (induct xq) (simp_all add: sup_assoc sup_commute sup_left_commute) |
|
864 |
||
30328 | 865 |
lemma sup_code [code]: |
866 |
"Seq f \<squnion> Seq g = Seq (\<lambda>u. case f () |
|
867 |
of Empty \<Rightarrow> g () |
|
868 |
| Insert x P \<Rightarrow> Insert x (P \<squnion> Seq g) |
|
30376 | 869 |
| Join P xq \<Rightarrow> adjunct (Seq g) (Join P xq))" |
30328 | 870 |
proof (cases "f ()") |
871 |
case Empty |
|
872 |
thus ?thesis |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
33988
diff
changeset
|
873 |
unfolding Seq_def by (simp add: sup_commute [of "\<bottom>"]) |
30328 | 874 |
next |
875 |
case Insert |
|
876 |
thus ?thesis |
|
877 |
unfolding Seq_def by (simp add: sup_assoc) |
|
878 |
next |
|
879 |
case Join |
|
880 |
thus ?thesis |
|
30376 | 881 |
unfolding Seq_def |
882 |
by (simp add: adjunct_sup sup_assoc sup_commute sup_left_commute) |
|
30328 | 883 |
qed |
884 |
||
30430
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
885 |
primrec contained :: "'a seq \<Rightarrow> 'a pred \<Rightarrow> bool" where |
44415 | 886 |
"contained Empty Q \<longleftrightarrow> True" |
887 |
| "contained (Insert x P) Q \<longleftrightarrow> eval Q x \<and> P \<le> Q" |
|
888 |
| "contained (Join P xq) Q \<longleftrightarrow> P \<le> Q \<and> contained xq Q" |
|
30430
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
889 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
890 |
lemma single_less_eq_eval: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
891 |
"single x \<le> P \<longleftrightarrow> eval P x" |
44415 | 892 |
by (auto simp add: less_eq_pred_def le_fun_def) |
30430
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
893 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
894 |
lemma contained_less_eq: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
895 |
"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
|
896 |
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
|
897 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
898 |
lemma less_eq_pred_code [code]: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
899 |
"Seq f \<le> Q = (case f () |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
900 |
of Empty \<Rightarrow> True |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
901 |
| 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
|
902 |
| 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
|
903 |
by (cases "f ()") |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
904 |
(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
|
905 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
906 |
lemma eq_pred_code [code]: |
31133 | 907 |
fixes P Q :: "'a pred" |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38651
diff
changeset
|
908 |
shows "HOL.equal P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P" |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38651
diff
changeset
|
909 |
by (auto simp add: equal) |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38651
diff
changeset
|
910 |
|
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38651
diff
changeset
|
911 |
lemma [code nbe]: |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38651
diff
changeset
|
912 |
"HOL.equal (x :: 'a pred) x \<longleftrightarrow> True" |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38651
diff
changeset
|
913 |
by (fact equal_refl) |
30430
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
914 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
915 |
lemma [code]: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
916 |
"pred_case f P = f (eval P)" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
917 |
by (cases P) simp |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
918 |
|
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
919 |
lemma [code]: |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
920 |
"pred_rec f P = f (eval P)" |
42ea5d85edcc
explicit code equations for some rarely used pred operations
haftmann
parents:
30378
diff
changeset
|
921 |
by (cases P) simp |
30328 | 922 |
|
31105
95f66b234086
added general preprocessing of equality in predicates for code generation
bulwahn
parents:
30430
diff
changeset
|
923 |
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
|
924 |
|
95f66b234086
added general preprocessing of equality in predicates for code generation
bulwahn
parents:
30430
diff
changeset
|
925 |
lemma eq_is_eq: "eq x y \<equiv> (x = y)" |
31108 | 926 |
by (rule eq_reflection) (auto intro: eq.intros elim: eq.cases) |
30948 | 927 |
|
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
928 |
primrec null :: "'a seq \<Rightarrow> bool" where |
44415 | 929 |
"null Empty \<longleftrightarrow> True" |
930 |
| "null (Insert x P) \<longleftrightarrow> False" |
|
931 |
| "null (Join P xq) \<longleftrightarrow> is_empty P \<and> null xq" |
|
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
932 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
933 |
lemma null_is_empty: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
934 |
"null xq \<longleftrightarrow> is_empty (pred_of_seq xq)" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
935 |
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
|
936 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
937 |
lemma is_empty_code [code]: |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
938 |
"is_empty (Seq f) \<longleftrightarrow> null (f ())" |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
939 |
by (simp add: null_is_empty Seq_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
940 |
|
33111 | 941 |
primrec the_only :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a seq \<Rightarrow> 'a" where |
942 |
[code del]: "the_only dfault Empty = dfault ()" |
|
44415 | 943 |
| "the_only dfault (Insert x P) = (if is_empty P then x else let y = singleton dfault P in if x = y then x else dfault ())" |
944 |
| "the_only dfault (Join P xq) = (if is_empty P then the_only dfault xq else if null xq then singleton dfault P |
|
33110 | 945 |
else let x = singleton dfault P; y = the_only dfault xq in |
33111 | 946 |
if x = y then x else dfault ())" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
947 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
948 |
lemma the_only_singleton: |
33110 | 949 |
"the_only dfault xq = singleton dfault (pred_of_seq xq)" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
950 |
by (induct xq) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
951 |
(auto simp add: singleton_bot singleton_single is_empty_def |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
952 |
null_is_empty Let_def singleton_sup) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
953 |
|
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
954 |
lemma singleton_code [code]: |
33110 | 955 |
"singleton dfault (Seq f) = (case f () |
33111 | 956 |
of Empty \<Rightarrow> dfault () |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
957 |
| Insert x P \<Rightarrow> if is_empty P then x |
33110 | 958 |
else let y = singleton dfault P in |
33111 | 959 |
if x = y then x else dfault () |
33110 | 960 |
| Join P xq \<Rightarrow> if is_empty P then the_only dfault xq |
961 |
else if null xq then singleton dfault P |
|
962 |
else let x = singleton dfault P; y = the_only dfault xq in |
|
33111 | 963 |
if x = y then x else dfault ())" |
32578
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
964 |
by (cases "f ()") |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
965 |
(auto simp add: Seq_def the_only_singleton is_empty_def |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
966 |
null_is_empty singleton_bot singleton_single singleton_sup Let_def) |
22117a76f943
added emptiness check predicate and singleton projection
haftmann
parents:
32372
diff
changeset
|
967 |
|
44414 | 968 |
definition the :: "'a pred \<Rightarrow> 'a" where |
37767 | 969 |
"the A = (THE x. eval A x)" |
33111 | 970 |
|
40674
54dbe6a1c349
adhere established Collect/mem convention more closely
haftmann
parents:
40616
diff
changeset
|
971 |
lemma the_eqI: |
41080 | 972 |
"(THE x. eval P x) = x \<Longrightarrow> the P = x" |
40674
54dbe6a1c349
adhere established Collect/mem convention more closely
haftmann
parents:
40616
diff
changeset
|
973 |
by (simp add: the_def) |
54dbe6a1c349
adhere established Collect/mem convention more closely
haftmann
parents:
40616
diff
changeset
|
974 |
|
44414 | 975 |
definition not_unique :: "'a pred \<Rightarrow> 'a" where |
976 |
[code del]: "not_unique A = (THE x. eval A x)" |
|
977 |
||
978 |
code_abort not_unique |
|
979 |
||
40674
54dbe6a1c349
adhere established Collect/mem convention more closely
haftmann
parents:
40616
diff
changeset
|
980 |
lemma the_eq [code]: "the A = singleton (\<lambda>x. not_unique A) A" |
54dbe6a1c349
adhere established Collect/mem convention more closely
haftmann
parents:
40616
diff
changeset
|
981 |
by (rule the_eqI) (simp add: singleton_def not_unique_def) |
33110 | 982 |
|
36531
19f6e3b0d9b6
code_reflect: specify module name directly after keyword
haftmann
parents:
36513
diff
changeset
|
983 |
code_reflect Predicate |
36513 | 984 |
datatypes pred = Seq and seq = Empty | Insert | Join |
985 |
functions map |
|
986 |
||
30948 | 987 |
ML {* |
988 |
signature PREDICATE = |
|
989 |
sig |
|
990 |
datatype 'a pred = Seq of (unit -> 'a seq) |
|
991 |
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
|
992 |
val yield: 'a pred -> ('a * 'a pred) option |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
993 |
val yieldn: int -> 'a pred -> 'a list * 'a pred |
31222 | 994 |
val map: ('a -> 'b) -> 'a pred -> 'b pred |
30948 | 995 |
end; |
996 |
||
997 |
structure Predicate : PREDICATE = |
|
998 |
struct |
|
999 |
||
36513 | 1000 |
datatype pred = datatype Predicate.pred |
1001 |
datatype seq = datatype Predicate.seq |
|
1002 |
||
1003 |
fun map f = Predicate.map f; |
|
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1004 |
|
36513 | 1005 |
fun yield (Seq f) = next (f ()) |
1006 |
and next Empty = NONE |
|
1007 |
| next (Insert (x, P)) = SOME (x, P) |
|
1008 |
| next (Join (P, xq)) = (case yield P |
|
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1009 |
of NONE => next xq |
36513 | 1010 |
| SOME (x, Q) => SOME (x, Seq (fn _ => Join (Q, xq)))); |
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1011 |
|
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1012 |
fun anamorph f k x = (if k = 0 then ([], x) |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1013 |
else case f x |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1014 |
of NONE => ([], x) |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1015 |
| SOME (v, y) => let |
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1016 |
val (vs, z) = anamorph f (k - 1) y |
33607 | 1017 |
in (v :: vs, z) end); |
30959
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1018 |
|
458e55fd0a33
fixed compilation of predicate types in ML environment
haftmann
parents:
30948
diff
changeset
|
1019 |
fun yieldn P = anamorph yield P; |
30948 | 1020 |
|
1021 |
end; |
|
1022 |
*} |
|
1023 |
||
46038
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1024 |
text {* Conversion from and to sets *} |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1025 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1026 |
definition pred_of_set :: "'a set \<Rightarrow> 'a pred" where |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1027 |
"pred_of_set = Pred \<circ> (\<lambda>A x. x \<in> A)" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1028 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1029 |
lemma eval_pred_of_set [simp]: |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1030 |
"eval (pred_of_set A) x \<longleftrightarrow> x \<in>A" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1031 |
by (simp add: pred_of_set_def) |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1032 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1033 |
definition set_of_pred :: "'a pred \<Rightarrow> 'a set" where |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1034 |
"set_of_pred = Collect \<circ> eval" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1035 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1036 |
lemma member_set_of_pred [simp]: |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1037 |
"x \<in> set_of_pred P \<longleftrightarrow> Predicate.eval P x" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1038 |
by (simp add: set_of_pred_def) |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1039 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1040 |
definition set_of_seq :: "'a seq \<Rightarrow> 'a set" where |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1041 |
"set_of_seq = set_of_pred \<circ> pred_of_seq" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1042 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1043 |
lemma member_set_of_seq [simp]: |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1044 |
"x \<in> set_of_seq xq = Predicate.member xq x" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1045 |
by (simp add: set_of_seq_def eval_member) |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1046 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1047 |
lemma of_pred_code [code]: |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1048 |
"set_of_pred (Predicate.Seq f) = (case f () of |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1049 |
Predicate.Empty \<Rightarrow> {} |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1050 |
| Predicate.Insert x P \<Rightarrow> insert x (set_of_pred P) |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1051 |
| Predicate.Join P xq \<Rightarrow> set_of_pred P \<union> set_of_seq xq)" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1052 |
by (auto split: seq.split simp add: eval_code) |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1053 |
|
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1054 |
lemma of_seq_code [code]: |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1055 |
"set_of_seq Predicate.Empty = {}" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1056 |
"set_of_seq (Predicate.Insert x P) = insert x (set_of_pred P)" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1057 |
"set_of_seq (Predicate.Join P xq) = set_of_pred P \<union> set_of_seq xq" |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1058 |
by auto |
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
haftmann
parents:
45970
diff
changeset
|
1059 |
|
30328 | 1060 |
no_notation |
41082 | 1061 |
bot ("\<bottom>") and |
1062 |
top ("\<top>") and |
|
30328 | 1063 |
inf (infixl "\<sqinter>" 70) and |
1064 |
sup (infixl "\<squnion>" 65) and |
|
1065 |
Inf ("\<Sqinter>_" [900] 900) and |
|
1066 |
Sup ("\<Squnion>_" [900] 900) and |
|
1067 |
bind (infixl "\<guillemotright>=" 70) |
|
1068 |
||
41080 | 1069 |
no_syntax (xsymbols) |
41082 | 1070 |
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10) |
1071 |
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10) |
|
41080 | 1072 |
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10) |
1073 |
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10) |
|
1074 |
||
36176
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents:
36008
diff
changeset
|
1075 |
hide_type (open) pred seq |
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents:
36008
diff
changeset
|
1076 |
hide_const (open) Pred eval single bind is_empty singleton if_pred not_pred holds |
33111 | 1077 |
Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map not_unique the |
30328 | 1078 |
|
1079 |
end |