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