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