| author | paulson <lp15@cam.ac.uk> | 
| Tue, 23 Apr 2024 10:26:04 +0100 | |
| changeset 80143 | 378593bf5109 | 
| parent 69861 | 62e47f06d22c | 
| child 80932 | 261cd8722677 | 
| permissions | -rw-r--r-- | 
| 
22259
 
476604be7d88
New theory for converting between predicates and sets.
 
berghofe 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Predicate.thy  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
2  | 
Author: 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  | 
|
| 60758 | 5  | 
section \<open>Predicates as enumerations\<close>  | 
| 
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  | 
| 53943 | 8  | 
imports String  | 
| 
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  | 
|
| 60758 | 11  | 
subsection \<open>The type of predicate enumerations (a monad)\<close>  | 
| 30328 | 12  | 
|
| 
66012
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
13  | 
datatype (plugins only: extraction) (dead 'a) pred = Pred (eval: "'a \<Rightarrow> bool")  | 
| 30328 | 14  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
15  | 
lemma pred_eqI:  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
16  | 
"(\<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
 | 
17  | 
by (cases P, cases Q) (auto simp add: fun_eq_iff)  | 
| 30328 | 18  | 
|
| 
46038
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
19  | 
lemma pred_eq_iff:  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
20  | 
"P = Q \<Longrightarrow> (\<And>w. eval P w \<longleftrightarrow> eval Q w)"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
21  | 
by (simp add: pred_eqI)  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
22  | 
|
| 44033 | 23  | 
instantiation pred :: (type) complete_lattice  | 
| 30328 | 24  | 
begin  | 
25  | 
||
26  | 
definition  | 
|
27  | 
"P \<le> Q \<longleftrightarrow> eval P \<le> eval Q"  | 
|
28  | 
||
29  | 
definition  | 
|
30  | 
"P < Q \<longleftrightarrow> eval P < eval Q"  | 
|
31  | 
||
32  | 
definition  | 
|
33  | 
"\<bottom> = Pred \<bottom>"  | 
|
34  | 
||
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
35  | 
lemma eval_bot [simp]:  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
36  | 
"eval \<bottom> = \<bottom>"  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
37  | 
by (simp add: bot_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
38  | 
|
| 30328 | 39  | 
definition  | 
40  | 
"\<top> = Pred \<top>"  | 
|
41  | 
||
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
42  | 
lemma eval_top [simp]:  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
43  | 
"eval \<top> = \<top>"  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
44  | 
by (simp add: top_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
45  | 
|
| 30328 | 46  | 
definition  | 
47  | 
"P \<sqinter> Q = Pred (eval P \<sqinter> eval Q)"  | 
|
48  | 
||
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
49  | 
lemma eval_inf [simp]:  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
50  | 
"eval (P \<sqinter> Q) = eval P \<sqinter> eval Q"  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
51  | 
by (simp add: inf_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
52  | 
|
| 30328 | 53  | 
definition  | 
54  | 
"P \<squnion> Q = Pred (eval P \<squnion> eval Q)"  | 
|
55  | 
||
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
56  | 
lemma eval_sup [simp]:  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
57  | 
"eval (P \<squnion> Q) = eval P \<squnion> eval Q"  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
58  | 
by (simp add: sup_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
59  | 
|
| 30328 | 60  | 
definition  | 
| 69275 | 61  | 
"\<Sqinter>A = Pred (\<Sqinter>(eval ` A))"  | 
| 30328 | 62  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
63  | 
lemma eval_Inf [simp]:  | 
| 69275 | 64  | 
"eval (\<Sqinter>A) = \<Sqinter>(eval ` A)"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
65  | 
by (simp add: Inf_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
66  | 
|
| 30328 | 67  | 
definition  | 
| 69275 | 68  | 
"\<Squnion>A = Pred (\<Squnion>(eval ` A))"  | 
| 30328 | 69  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
70  | 
lemma eval_Sup [simp]:  | 
| 69275 | 71  | 
"eval (\<Squnion>A) = \<Squnion>(eval ` A)"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
72  | 
by (simp add: Sup_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
73  | 
|
| 44033 | 74  | 
instance proof  | 
| 44415 | 75  | 
qed (auto intro!: pred_eqI simp add: less_eq_pred_def less_pred_def le_fun_def less_fun_def)  | 
| 44033 | 76  | 
|
77  | 
end  | 
|
78  | 
||
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
79  | 
lemma eval_INF [simp]:  | 
| 69275 | 80  | 
"eval (\<Sqinter>(f ` A)) = \<Sqinter>((eval \<circ> f) ` A)"  | 
| 
69861
 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 
haftmann 
parents: 
69275 
diff
changeset
 | 
81  | 
by (simp add: image_comp)  | 
| 44033 | 82  | 
|
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
83  | 
lemma eval_SUP [simp]:  | 
| 69275 | 84  | 
"eval (\<Squnion>(f ` A)) = \<Squnion>((eval \<circ> f) ` A)"  | 
| 
69861
 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 
haftmann 
parents: 
69275 
diff
changeset
 | 
85  | 
by (simp add: image_comp)  | 
| 44033 | 86  | 
|
87  | 
instantiation pred :: (type) complete_boolean_algebra  | 
|
88  | 
begin  | 
|
89  | 
||
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
90  | 
definition  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
91  | 
"- P = Pred (- eval P)"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
92  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
93  | 
lemma eval_compl [simp]:  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
94  | 
"eval (- P) = - eval P"  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
95  | 
by (simp add: uminus_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
96  | 
|
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
97  | 
definition  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
98  | 
"P - Q = Pred (eval P - eval Q)"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
99  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
100  | 
lemma eval_minus [simp]:  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
101  | 
"eval (P - Q) = eval P - eval Q"  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
102  | 
by (simp add: minus_pred_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
103  | 
|
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
104  | 
instance proof  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
105  | 
fix A::"'a pred set set"  | 
| 69275 | 106  | 
  show "\<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
 | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
107  | 
proof (simp add: less_eq_pred_def Sup_fun_def Inf_fun_def, safe)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
108  | 
fix w  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
109  | 
assume A: "\<forall>x\<in>A. \<exists>f\<in>x. eval f w"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
110  | 
define F where "F = (\<lambda> x . SOME f . f \<in> x \<and> eval f w)"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
111  | 
have [simp]: "(\<forall>f\<in> (F ` A). eval f w)"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
112  | 
by (metis (no_types, lifting) A F_def image_iff some_eq_ex)  | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
113  | 
have "(\<exists>f. F ` A = f ` A \<and> (\<forall>Y\<in>A. f Y \<in> Y)) \<and> (\<forall>f\<in>(F ` A). eval f w)"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
114  | 
using A by (simp, metis (no_types, lifting) F_def someI)+  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
115  | 
from this show "\<exists>x. (\<exists>f. x = f ` A \<and> (\<forall>Y\<in>A. f Y \<in> Y)) \<and> (\<forall>f\<in>x. eval f w)"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
116  | 
by (rule exI [of _ "F ` A"])  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67399 
diff
changeset
 | 
117  | 
qed  | 
| 46884 | 118  | 
qed (auto intro!: pred_eqI)  | 
| 30328 | 119  | 
|
| 
22259
 
476604be7d88
New theory for converting between predicates and sets.
 
berghofe 
parents:  
diff
changeset
 | 
120  | 
end  | 
| 30328 | 121  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
122  | 
definition single :: "'a \<Rightarrow> 'a pred" where  | 
| 67399 | 123  | 
"single x = Pred ((=) x)"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
124  | 
|
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
125  | 
lemma eval_single [simp]:  | 
| 67399 | 126  | 
"eval (single x) = (=) x"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
127  | 
by (simp add: single_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
128  | 
|
| 62026 | 129  | 
definition bind :: "'a pred \<Rightarrow> ('a \<Rightarrow> 'b pred) \<Rightarrow> 'b pred" (infixl "\<bind>" 70) where
 | 
| 69275 | 130  | 
  "P \<bind> f = (\<Squnion>(f ` {x. eval P x}))"
 | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
131  | 
|
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
132  | 
lemma eval_bind [simp]:  | 
| 69275 | 133  | 
  "eval (P \<bind> f) = eval (\<Squnion>(f ` {x. eval P x}))"
 | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
134  | 
by (simp add: bind_def)  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
135  | 
|
| 30328 | 136  | 
lemma bind_bind:  | 
| 62026 | 137  | 
"(P \<bind> Q) \<bind> R = P \<bind> (\<lambda>x. Q x \<bind> R)"  | 
| 46884 | 138  | 
by (rule pred_eqI) auto  | 
| 30328 | 139  | 
|
140  | 
lemma bind_single:  | 
|
| 62026 | 141  | 
"P \<bind> single = P"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
142  | 
by (rule pred_eqI) auto  | 
| 30328 | 143  | 
|
144  | 
lemma single_bind:  | 
|
| 62026 | 145  | 
"single x \<bind> P = P x"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
146  | 
by (rule pred_eqI) auto  | 
| 30328 | 147  | 
|
148  | 
lemma bottom_bind:  | 
|
| 62026 | 149  | 
"\<bottom> \<bind> P = \<bottom>"  | 
| 
40674
 
54dbe6a1c349
adhere established Collect/mem convention more closely
 
haftmann 
parents: 
40616 
diff
changeset
 | 
150  | 
by (rule pred_eqI) auto  | 
| 30328 | 151  | 
|
152  | 
lemma sup_bind:  | 
|
| 62026 | 153  | 
"(P \<squnion> Q) \<bind> R = P \<bind> R \<squnion> Q \<bind> R"  | 
| 
40674
 
54dbe6a1c349
adhere established Collect/mem convention more closely
 
haftmann 
parents: 
40616 
diff
changeset
 | 
154  | 
by (rule pred_eqI) auto  | 
| 30328 | 155  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
156  | 
lemma Sup_bind:  | 
| 62026 | 157  | 
"(\<Squnion>A \<bind> f) = \<Squnion>((\<lambda>x. x \<bind> f) ` A)"  | 
| 46884 | 158  | 
by (rule pred_eqI) auto  | 
| 30328 | 159  | 
|
160  | 
lemma pred_iffI:  | 
|
161  | 
assumes "\<And>x. eval A x \<Longrightarrow> eval B x"  | 
|
162  | 
and "\<And>x. eval B x \<Longrightarrow> eval A x"  | 
|
163  | 
shows "A = B"  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
164  | 
using assms by (auto intro: pred_eqI)  | 
| 30328 | 165  | 
|
166  | 
lemma singleI: "eval (single x) x"  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
167  | 
by simp  | 
| 30328 | 168  | 
|
169  | 
lemma singleI_unit: "eval (single ()) x"  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
170  | 
by simp  | 
| 30328 | 171  | 
|
172  | 
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
 | 
173  | 
by simp  | 
| 30328 | 174  | 
|
175  | 
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
 | 
176  | 
by simp  | 
| 30328 | 177  | 
|
| 62026 | 178  | 
lemma bindI: "eval P x \<Longrightarrow> eval (Q x) y \<Longrightarrow> eval (P \<bind> Q) y"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
179  | 
by auto  | 
| 30328 | 180  | 
|
| 62026 | 181  | 
lemma bindE: "eval (R \<bind> 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
 | 
182  | 
by auto  | 
| 30328 | 183  | 
|
184  | 
lemma botE: "eval \<bottom> x \<Longrightarrow> P"  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
185  | 
by auto  | 
| 30328 | 186  | 
|
187  | 
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
 | 
188  | 
by auto  | 
| 30328 | 189  | 
|
190  | 
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
 | 
191  | 
by auto  | 
| 30328 | 192  | 
|
193  | 
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
 | 
194  | 
by auto  | 
| 30328 | 195  | 
|
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
196  | 
lemma single_not_bot [simp]:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
197  | 
"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
 | 
198  | 
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
 | 
199  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
200  | 
lemma not_bot:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
201  | 
assumes "A \<noteq> \<bottom>"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
202  | 
obtains x where "eval A x"  | 
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
45630 
diff
changeset
 | 
203  | 
using assms by (cases A) (auto simp add: bot_pred_def)  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
45630 
diff
changeset
 | 
204  | 
|
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
205  | 
|
| 60758 | 206  | 
subsection \<open>Emptiness check and definite choice\<close>  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
207  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
208  | 
definition is_empty :: "'a pred \<Rightarrow> bool" where  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
209  | 
"is_empty A \<longleftrightarrow> A = \<bottom>"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
210  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
211  | 
lemma is_empty_bot:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
212  | 
"is_empty \<bottom>"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
213  | 
by (simp add: is_empty_def)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
214  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
215  | 
lemma not_is_empty_single:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
216  | 
"\<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
 | 
217  | 
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
 | 
218  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
219  | 
lemma is_empty_sup:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
220  | 
"is_empty (A \<squnion> B) \<longleftrightarrow> is_empty A \<and> is_empty B"  | 
| 36008 | 221  | 
by (auto simp add: is_empty_def)  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
222  | 
|
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
223  | 
definition singleton :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a pred \<Rightarrow> 'a" where  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
224  | 
"singleton default A = (if \<exists>!x. eval A x then THE x. eval A x else default ())" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
225  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
226  | 
lemma singleton_eqI:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
227  | 
"\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton default A = x" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
228  | 
by (auto simp add: singleton_def)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
229  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
230  | 
lemma eval_singletonI:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
231  | 
"\<exists>!x. eval A x \<Longrightarrow> eval A (singleton default A)" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
232  | 
proof -  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
233  | 
assume assm: "\<exists>!x. eval A x"  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51143 
diff
changeset
 | 
234  | 
then obtain x where x: "eval A x" ..  | 
| 60166 | 235  | 
with assm have "singleton default A = x" by (rule singleton_eqI)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51143 
diff
changeset
 | 
236  | 
with x show ?thesis by simp  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
237  | 
qed  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
238  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
239  | 
lemma single_singleton:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
240  | 
"\<exists>!x. eval A x \<Longrightarrow> single (singleton default A) = A" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
241  | 
proof -  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
242  | 
assume assm: "\<exists>!x. eval A x"  | 
| 60166 | 243  | 
then have "eval A (singleton default A)"  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
244  | 
by (rule eval_singletonI)  | 
| 60166 | 245  | 
moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton default A = x"  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
246  | 
by (rule singleton_eqI)  | 
| 60166 | 247  | 
ultimately have "eval (single (singleton default 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
 | 
248  | 
by (simp (no_asm_use) add: single_def fun_eq_iff) blast  | 
| 60166 | 249  | 
then have "\<And>x. eval (single (singleton default A)) x = eval A x"  | 
| 
40616
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
250  | 
by simp  | 
| 
 
c5ee1e06d795
eval simp rules for predicate type, simplify primitive proofs
 
haftmann 
parents: 
39302 
diff
changeset
 | 
251  | 
then show ?thesis by (rule pred_eqI)  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
252  | 
qed  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
253  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
254  | 
lemma singleton_undefinedI:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
255  | 
"\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton default A = default ()" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
256  | 
by (simp add: singleton_def)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
257  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
258  | 
lemma singleton_bot:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
259  | 
"singleton default \<bottom> = default ()" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
260  | 
by (auto simp add: bot_pred_def intro: singleton_undefinedI)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
261  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
262  | 
lemma singleton_single:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
263  | 
"singleton default (single x) = x" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
264  | 
by (auto simp add: intro: singleton_eqI singleI elim: singleE)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
265  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
266  | 
lemma singleton_sup_single_single:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
267  | 
"singleton default (single x \<squnion> single y) = (if x = y then x else default ())" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
268  | 
proof (cases "x = y")  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
269  | 
case True then show ?thesis by (simp add: singleton_single)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
270  | 
next  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
271  | 
case False  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
272  | 
have "eval (single x \<squnion> single y) x"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
273  | 
and "eval (single x \<squnion> single y) y"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
274  | 
by (auto intro: supI1 supI2 singleI)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
275  | 
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
 | 
276  | 
by blast  | 
| 60166 | 277  | 
then have "singleton default (single x \<squnion> single y) = default ()"  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
278  | 
by (rule singleton_undefinedI)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
279  | 
with False show ?thesis by simp  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
280  | 
qed  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
281  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
282  | 
lemma singleton_sup_aux:  | 
| 60166 | 283  | 
"singleton default (A \<squnion> B) = (if A = \<bottom> then singleton default B  | 
284  | 
else if B = \<bottom> then singleton default A  | 
|
285  | 
else singleton default  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
286  | 
(single (singleton default A) \<squnion> single (singleton default B)))" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
287  | 
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
 | 
288  | 
case True then show ?thesis by (simp add: single_singleton)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
289  | 
next  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
290  | 
case False  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
291  | 
from False have A_or_B:  | 
| 60166 | 292  | 
"singleton default A = default () \<or> singleton default B = default ()"  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
293  | 
by (auto intro!: singleton_undefinedI)  | 
| 60166 | 294  | 
then have rhs: "singleton default  | 
295  | 
(single (singleton default A) \<squnion> single (singleton default B)) = default ()"  | 
|
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
296  | 
by (auto simp add: singleton_sup_single_single singleton_single)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
297  | 
from False have not_unique:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
298  | 
"\<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
 | 
299  | 
show ?thesis proof (cases "A \<noteq> \<bottom> \<and> B \<noteq> \<bottom>")  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
300  | 
case True  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
301  | 
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
 | 
302  | 
by (blast elim: not_bot)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
303  | 
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
 | 
304  | 
by (auto simp add: sup_pred_def bot_pred_def)  | 
| 60166 | 305  | 
then have "singleton default (A \<squnion> B) = default ()" by (rule singleton_undefinedI)  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
306  | 
with True rhs show ?thesis by simp  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
307  | 
next  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
308  | 
case False then show ?thesis by auto  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
309  | 
qed  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
310  | 
qed  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
311  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
312  | 
lemma singleton_sup:  | 
| 60166 | 313  | 
"singleton default (A \<squnion> B) = (if A = \<bottom> then singleton default B  | 
314  | 
else if B = \<bottom> then singleton default A  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
315  | 
else if singleton default A = singleton default B then singleton default A else default ())" for default  | 
| 60166 | 316  | 
using singleton_sup_aux [of default A B] by (simp only: singleton_sup_single_single)  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
317  | 
|
| 30328 | 318  | 
|
| 60758 | 319  | 
subsection \<open>Derived operations\<close>  | 
| 30328 | 320  | 
|
321  | 
definition if_pred :: "bool \<Rightarrow> unit pred" where  | 
|
322  | 
if_pred_eq: "if_pred b = (if b then single () else \<bottom>)"  | 
|
323  | 
||
| 
33754
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
324  | 
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
 | 
325  | 
holds_eq: "holds P = eval P ()"  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
326  | 
|
| 30328 | 327  | 
definition not_pred :: "unit pred \<Rightarrow> unit pred" where  | 
328  | 
not_pred_eq: "not_pred P = (if eval P () then \<bottom> else single ())"  | 
|
329  | 
||
330  | 
lemma if_predI: "P \<Longrightarrow> eval (if_pred P) ()"  | 
|
331  | 
unfolding if_pred_eq by (auto intro: singleI)  | 
|
332  | 
||
333  | 
lemma if_predE: "eval (if_pred b) x \<Longrightarrow> (b \<Longrightarrow> x = () \<Longrightarrow> P) \<Longrightarrow> P"  | 
|
334  | 
unfolding if_pred_eq by (cases b) (auto elim: botE)  | 
|
335  | 
||
336  | 
lemma not_predI: "\<not> P \<Longrightarrow> eval (not_pred (Pred (\<lambda>u. P))) ()"  | 
|
| 
66012
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
337  | 
unfolding not_pred_eq by (auto intro: singleI)  | 
| 30328 | 338  | 
|
339  | 
lemma not_predI': "\<not> eval P () \<Longrightarrow> eval (not_pred P) ()"  | 
|
340  | 
unfolding not_pred_eq by (auto intro: singleI)  | 
|
341  | 
||
342  | 
lemma not_predE: "eval (not_pred (Pred (\<lambda>u. P))) x \<Longrightarrow> (\<not> P \<Longrightarrow> thesis) \<Longrightarrow> thesis"  | 
|
343  | 
unfolding not_pred_eq  | 
|
| 62390 | 344  | 
by (auto split: if_split_asm elim: botE)  | 
| 30328 | 345  | 
|
346  | 
lemma not_predE': "eval (not_pred P) x \<Longrightarrow> (\<not> eval P x \<Longrightarrow> thesis) \<Longrightarrow> thesis"  | 
|
347  | 
unfolding not_pred_eq  | 
|
| 62390 | 348  | 
by (auto split: if_split_asm elim: botE)  | 
| 
33754
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
349  | 
lemma "f () = False \<or> f () = True"  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
350  | 
by simp  | 
| 30328 | 351  | 
|
| 37549 | 352  | 
lemma closure_of_bool_cases [no_atp]:  | 
| 44007 | 353  | 
fixes f :: "unit \<Rightarrow> bool"  | 
354  | 
assumes "f = (\<lambda>u. False) \<Longrightarrow> P f"  | 
|
355  | 
assumes "f = (\<lambda>u. True) \<Longrightarrow> P f"  | 
|
356  | 
shows "P f"  | 
|
| 
33754
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
357  | 
proof -  | 
| 44007 | 358  | 
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
 | 
359  | 
apply (cases "f ()")  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
360  | 
apply (rule disjI2)  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
361  | 
apply (rule ext)  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
362  | 
apply (simp add: unit_eq)  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
363  | 
apply (rule disjI1)  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
364  | 
apply (rule ext)  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
365  | 
apply (simp add: unit_eq)  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
366  | 
done  | 
| 41550 | 367  | 
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
 | 
368  | 
qed  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
369  | 
|
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
370  | 
lemma unit_pred_cases:  | 
| 44007 | 371  | 
assumes "P \<bottom>"  | 
372  | 
assumes "P (single ())"  | 
|
373  | 
shows "P Q"  | 
|
| 44415 | 374  | 
using assms unfolding bot_pred_def bot_fun_def bot_bool_def empty_def single_def proof (cases Q)  | 
| 44007 | 375  | 
fix f  | 
376  | 
assume "P (Pred (\<lambda>u. False))" "P (Pred (\<lambda>u. () = u))"  | 
|
377  | 
then have "P (Pred f)"  | 
|
378  | 
by (cases _ f rule: closure_of_bool_cases) simp_all  | 
|
379  | 
moreover assume "Q = Pred f"  | 
|
380  | 
ultimately show "P Q" by simp  | 
|
381  | 
qed  | 
|
382  | 
||
| 
33754
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
383  | 
lemma holds_if_pred:  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
384  | 
"holds (if_pred b) = b"  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
385  | 
unfolding if_pred_eq holds_eq  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
386  | 
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
 | 
387  | 
|
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
388  | 
lemma if_pred_holds:  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
389  | 
"if_pred (holds P) = P"  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
390  | 
unfolding if_pred_eq holds_eq  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
391  | 
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
 | 
392  | 
|
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
393  | 
lemma is_empty_holds:  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
394  | 
"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
 | 
395  | 
unfolding is_empty_def holds_eq  | 
| 
 
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
 
bulwahn 
parents: 
33622 
diff
changeset
 | 
396  | 
by (rule unit_pred_cases) (auto elim: botE intro: singleI)  | 
| 30328 | 397  | 
|
| 41311 | 398  | 
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pred \<Rightarrow> 'b pred" where
 | 
| 67091 | 399  | 
"map f P = P \<bind> (single \<circ> f)"  | 
| 41311 | 400  | 
|
401  | 
lemma eval_map [simp]:  | 
|
| 44363 | 402  | 
  "eval (map f P) = (\<Squnion>x\<in>{x. eval P x}. (\<lambda>y. f x = y))"
 | 
| 
69861
 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 
haftmann 
parents: 
69275 
diff
changeset
 | 
403  | 
by (simp add: map_def comp_def image_comp)  | 
| 41311 | 404  | 
|
| 
55467
 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 
blanchet 
parents: 
55416 
diff
changeset
 | 
405  | 
functor map: map  | 
| 44363 | 406  | 
by (rule ext, rule pred_eqI, auto)+  | 
| 41311 | 407  | 
|
408  | 
||
| 60758 | 409  | 
subsection \<open>Implementation\<close>  | 
| 30328 | 410  | 
|
| 
58350
 
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
 
blanchet 
parents: 
58334 
diff
changeset
 | 
411  | 
datatype (plugins only: code extraction) (dead 'a) seq =  | 
| 58334 | 412  | 
Empty  | 
413  | 
| Insert "'a" "'a pred"  | 
|
414  | 
| Join "'a pred" "'a seq"  | 
|
| 30328 | 415  | 
|
416  | 
primrec pred_of_seq :: "'a seq \<Rightarrow> 'a pred" where  | 
|
| 44414 | 417  | 
"pred_of_seq Empty = \<bottom>"  | 
418  | 
| "pred_of_seq (Insert x P) = single x \<squnion> P"  | 
|
419  | 
| "pred_of_seq (Join P xq) = P \<squnion> pred_of_seq xq"  | 
|
| 30328 | 420  | 
|
421  | 
definition Seq :: "(unit \<Rightarrow> 'a seq) \<Rightarrow> 'a pred" where  | 
|
422  | 
"Seq f = pred_of_seq (f ())"  | 
|
423  | 
||
424  | 
code_datatype Seq  | 
|
425  | 
||
426  | 
primrec member :: "'a seq \<Rightarrow> 'a \<Rightarrow> bool" where  | 
|
427  | 
"member Empty x \<longleftrightarrow> False"  | 
|
| 44414 | 428  | 
| "member (Insert y P) x \<longleftrightarrow> x = y \<or> eval P x"  | 
429  | 
| "member (Join P xq) x \<longleftrightarrow> eval P x \<or> member xq x"  | 
|
| 30328 | 430  | 
|
431  | 
lemma eval_member:  | 
|
432  | 
"member xq = eval (pred_of_seq xq)"  | 
|
433  | 
proof (induct xq)  | 
|
434  | 
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
 | 
435  | 
by (auto simp add: fun_eq_iff elim: botE)  | 
| 30328 | 436  | 
next  | 
437  | 
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
 | 
438  | 
by (auto simp add: fun_eq_iff elim: supE singleE intro: supI1 supI2 singleI)  | 
| 30328 | 439  | 
next  | 
440  | 
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
 | 
441  | 
by (auto simp add: fun_eq_iff elim: supE intro: supI1 supI2)  | 
| 30328 | 442  | 
qed  | 
443  | 
||
| 
46038
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
444  | 
lemma eval_code [(* FIXME declare simp *)code]: "eval (Seq f) = member (f ())"  | 
| 30328 | 445  | 
unfolding Seq_def by (rule sym, rule eval_member)  | 
446  | 
||
447  | 
lemma single_code [code]:  | 
|
448  | 
"single x = Seq (\<lambda>u. Insert x \<bottom>)"  | 
|
449  | 
unfolding Seq_def by simp  | 
|
450  | 
||
| 41080 | 451  | 
primrec "apply" :: "('a \<Rightarrow> 'b pred) \<Rightarrow> 'a seq \<Rightarrow> 'b seq" where
 | 
| 44415 | 452  | 
"apply f Empty = Empty"  | 
| 62026 | 453  | 
| "apply f (Insert x P) = Join (f x) (Join (P \<bind> f) Empty)"  | 
454  | 
| "apply f (Join P xq) = Join (P \<bind> f) (apply f xq)"  | 
|
| 30328 | 455  | 
|
456  | 
lemma apply_bind:  | 
|
| 62026 | 457  | 
"pred_of_seq (apply f xq) = pred_of_seq xq \<bind> f"  | 
| 30328 | 458  | 
proof (induct xq)  | 
459  | 
case Empty show ?case  | 
|
460  | 
by (simp add: bottom_bind)  | 
|
461  | 
next  | 
|
462  | 
case Insert show ?case  | 
|
463  | 
by (simp add: single_bind sup_bind)  | 
|
464  | 
next  | 
|
465  | 
case Join then show ?case  | 
|
466  | 
by (simp add: sup_bind)  | 
|
467  | 
qed  | 
|
468  | 
||
469  | 
lemma bind_code [code]:  | 
|
| 62026 | 470  | 
"Seq g \<bind> f = Seq (\<lambda>u. apply f (g ()))"  | 
| 30328 | 471  | 
unfolding Seq_def by (rule sym, rule apply_bind)  | 
472  | 
||
473  | 
lemma bot_set_code [code]:  | 
|
474  | 
"\<bottom> = Seq (\<lambda>u. Empty)"  | 
|
475  | 
unfolding Seq_def by simp  | 
|
476  | 
||
| 30376 | 477  | 
primrec adjunct :: "'a pred \<Rightarrow> 'a seq \<Rightarrow> 'a seq" where  | 
| 44415 | 478  | 
"adjunct P Empty = Join P Empty"  | 
479  | 
| "adjunct P (Insert x Q) = Insert x (Q \<squnion> P)"  | 
|
480  | 
| "adjunct P (Join Q xq) = Join Q (adjunct P xq)"  | 
|
| 30376 | 481  | 
|
482  | 
lemma adjunct_sup:  | 
|
483  | 
"pred_of_seq (adjunct P xq) = P \<squnion> pred_of_seq xq"  | 
|
484  | 
by (induct xq) (simp_all add: sup_assoc sup_commute sup_left_commute)  | 
|
485  | 
||
| 30328 | 486  | 
lemma sup_code [code]:  | 
487  | 
"Seq f \<squnion> Seq g = Seq (\<lambda>u. case f ()  | 
|
488  | 
of Empty \<Rightarrow> g ()  | 
|
489  | 
| Insert x P \<Rightarrow> Insert x (P \<squnion> Seq g)  | 
|
| 30376 | 490  | 
| Join P xq \<Rightarrow> adjunct (Seq g) (Join P xq))"  | 
| 30328 | 491  | 
proof (cases "f ()")  | 
492  | 
case Empty  | 
|
493  | 
thus ?thesis  | 
|
| 
34007
 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 
haftmann 
parents: 
33988 
diff
changeset
 | 
494  | 
unfolding Seq_def by (simp add: sup_commute [of "\<bottom>"])  | 
| 30328 | 495  | 
next  | 
496  | 
case Insert  | 
|
497  | 
thus ?thesis  | 
|
498  | 
unfolding Seq_def by (simp add: sup_assoc)  | 
|
499  | 
next  | 
|
500  | 
case Join  | 
|
501  | 
thus ?thesis  | 
|
| 30376 | 502  | 
unfolding Seq_def  | 
503  | 
by (simp add: adjunct_sup sup_assoc sup_commute sup_left_commute)  | 
|
| 30328 | 504  | 
qed  | 
505  | 
||
| 
30430
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
506  | 
primrec contained :: "'a seq \<Rightarrow> 'a pred \<Rightarrow> bool" where  | 
| 44415 | 507  | 
"contained Empty Q \<longleftrightarrow> True"  | 
508  | 
| "contained (Insert x P) Q \<longleftrightarrow> eval Q x \<and> P \<le> Q"  | 
|
509  | 
| "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
 | 
510  | 
|
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
511  | 
lemma single_less_eq_eval:  | 
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
512  | 
"single x \<le> P \<longleftrightarrow> eval P x"  | 
| 44415 | 513  | 
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
 | 
514  | 
|
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
515  | 
lemma contained_less_eq:  | 
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
516  | 
"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
 | 
517  | 
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
 | 
518  | 
|
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
519  | 
lemma less_eq_pred_code [code]:  | 
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
520  | 
"Seq f \<le> Q = (case f ()  | 
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
521  | 
of Empty \<Rightarrow> True  | 
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
522  | 
| 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
 | 
523  | 
| 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
 | 
524  | 
by (cases "f ()")  | 
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
525  | 
(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
 | 
526  | 
|
| 
66012
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
527  | 
instantiation pred :: (type) equal  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
528  | 
begin  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
529  | 
|
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
530  | 
definition equal_pred  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
531  | 
where [simp]: "HOL.equal P Q \<longleftrightarrow> P = (Q :: 'a pred)"  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
532  | 
|
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
533  | 
instance by standard simp  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
534  | 
|
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
535  | 
end  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
536  | 
|
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
537  | 
lemma [code]:  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
538  | 
"HOL.equal P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P" for P Q :: "'a pred"  | 
| 
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
539  | 
by auto  | 
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38651 
diff
changeset
 | 
540  | 
|
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38651 
diff
changeset
 | 
541  | 
lemma [code nbe]:  | 
| 
66012
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
542  | 
"HOL.equal P P \<longleftrightarrow> True" for P :: "'a pred"  | 
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38651 
diff
changeset
 | 
543  | 
by (fact equal_refl)  | 
| 
30430
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
544  | 
|
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
545  | 
lemma [code]:  | 
| 55416 | 546  | 
"case_pred f P = f (eval P)"  | 
| 
66012
 
59bf29d2b3a1
modernized (code) setup for enumeration predicates
 
haftmann 
parents: 
62390 
diff
changeset
 | 
547  | 
by (fact pred.case_eq_if)  | 
| 
30430
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
548  | 
|
| 
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
549  | 
lemma [code]:  | 
| 55416 | 550  | 
"rec_pred f P = f (eval P)"  | 
| 
30430
 
42ea5d85edcc
explicit code equations for some rarely used pred operations
 
haftmann 
parents: 
30378 
diff
changeset
 | 
551  | 
by (cases P) simp  | 
| 30328 | 552  | 
|
| 
31105
 
95f66b234086
added general preprocessing of equality in predicates for code generation
 
bulwahn 
parents: 
30430 
diff
changeset
 | 
553  | 
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
 | 
554  | 
|
| 
 
95f66b234086
added general preprocessing of equality in predicates for code generation
 
bulwahn 
parents: 
30430 
diff
changeset
 | 
555  | 
lemma eq_is_eq: "eq x y \<equiv> (x = y)"  | 
| 31108 | 556  | 
by (rule eq_reflection) (auto intro: eq.intros elim: eq.cases)  | 
| 30948 | 557  | 
|
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
558  | 
primrec null :: "'a seq \<Rightarrow> bool" where  | 
| 44415 | 559  | 
"null Empty \<longleftrightarrow> True"  | 
560  | 
| "null (Insert x P) \<longleftrightarrow> False"  | 
|
561  | 
| "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
 | 
562  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
563  | 
lemma null_is_empty:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
564  | 
"null xq \<longleftrightarrow> is_empty (pred_of_seq xq)"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
565  | 
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
 | 
566  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
567  | 
lemma is_empty_code [code]:  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
568  | 
"is_empty (Seq f) \<longleftrightarrow> null (f ())"  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
569  | 
by (simp add: null_is_empty Seq_def)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
570  | 
|
| 33111 | 571  | 
primrec the_only :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a seq \<Rightarrow> 'a" where  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
572  | 
"the_only default Empty = default ()" for default  | 
| 
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
573  | 
| "the_only default (Insert x P) =  | 
| 
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
574  | 
(if is_empty P then x else let y = singleton default P in if x = y then x else default ())" for default  | 
| 
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
575  | 
| "the_only default (Join P xq) =  | 
| 60166 | 576  | 
(if is_empty P then the_only default xq else if null xq then singleton default P  | 
577  | 
else let x = singleton default P; y = the_only default xq in  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
578  | 
if x = y then x else default ())" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
579  | 
|
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
580  | 
lemma the_only_singleton:  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
581  | 
"the_only default xq = singleton default (pred_of_seq xq)" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
582  | 
by (induct xq)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
583  | 
(auto simp add: singleton_bot singleton_single is_empty_def  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
584  | 
null_is_empty Let_def singleton_sup)  | 
| 
 
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  | 
lemma singleton_code [code]:  | 
| 60166 | 587  | 
"singleton default (Seq f) =  | 
588  | 
(case f () of  | 
|
589  | 
Empty \<Rightarrow> default ()  | 
|
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
590  | 
| Insert x P \<Rightarrow> if is_empty P then x  | 
| 60166 | 591  | 
else let y = singleton default P in  | 
592  | 
if x = y then x else default ()  | 
|
593  | 
| Join P xq \<Rightarrow> if is_empty P then the_only default xq  | 
|
594  | 
else if null xq then singleton default P  | 
|
595  | 
else let x = singleton default P; y = the_only default xq in  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66012 
diff
changeset
 | 
596  | 
if x = y then x else default ())" for default  | 
| 
32578
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
597  | 
by (cases "f ()")  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
598  | 
(auto simp add: Seq_def the_only_singleton is_empty_def  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
599  | 
null_is_empty singleton_bot singleton_single singleton_sup Let_def)  | 
| 
 
22117a76f943
added emptiness check predicate and singleton projection
 
haftmann 
parents: 
32372 
diff
changeset
 | 
600  | 
|
| 44414 | 601  | 
definition the :: "'a pred \<Rightarrow> 'a" where  | 
| 37767 | 602  | 
"the A = (THE x. eval A x)"  | 
| 33111 | 603  | 
|
| 
40674
 
54dbe6a1c349
adhere established Collect/mem convention more closely
 
haftmann 
parents: 
40616 
diff
changeset
 | 
604  | 
lemma the_eqI:  | 
| 41080 | 605  | 
"(THE x. eval P x) = x \<Longrightarrow> the P = x"  | 
| 
40674
 
54dbe6a1c349
adhere established Collect/mem convention more closely
 
haftmann 
parents: 
40616 
diff
changeset
 | 
606  | 
by (simp add: the_def)  | 
| 
 
54dbe6a1c349
adhere established Collect/mem convention more closely
 
haftmann 
parents: 
40616 
diff
changeset
 | 
607  | 
|
| 53943 | 608  | 
lemma the_eq [code]: "the A = singleton (\<lambda>x. Code.abort (STR ''not_unique'') (\<lambda>_. the A)) A"  | 
609  | 
by (rule the_eqI) (simp add: singleton_def the_def)  | 
|
| 33110 | 610  | 
|
| 
36531
 
19f6e3b0d9b6
code_reflect: specify module name directly after keyword
 
haftmann 
parents: 
36513 
diff
changeset
 | 
611  | 
code_reflect Predicate  | 
| 36513 | 612  | 
datatypes pred = Seq and seq = Empty | Insert | Join  | 
613  | 
||
| 60758 | 614  | 
ML \<open>  | 
| 30948 | 615  | 
signature PREDICATE =  | 
616  | 
sig  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
617  | 
  val anamorph: ('a -> ('b * 'a) option) -> int -> 'a -> 'b list * 'a
 | 
| 30948 | 618  | 
datatype 'a pred = Seq of (unit -> 'a seq)  | 
619  | 
and 'a seq = Empty | Insert of 'a * 'a pred | Join of 'a pred * 'a seq  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
620  | 
  val map: ('a -> 'b) -> 'a pred -> 'b pred
 | 
| 
30959
 
458e55fd0a33
fixed compilation of predicate types in ML environment
 
haftmann 
parents: 
30948 
diff
changeset
 | 
621  | 
  val yield: 'a pred -> ('a * 'a pred) option
 | 
| 
 
458e55fd0a33
fixed compilation of predicate types in ML environment
 
haftmann 
parents: 
30948 
diff
changeset
 | 
622  | 
val yieldn: int -> 'a pred -> 'a list * 'a pred  | 
| 30948 | 623  | 
end;  | 
624  | 
||
625  | 
structure Predicate : PREDICATE =  | 
|
626  | 
struct  | 
|
627  | 
||
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
628  | 
fun anamorph f k x =  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
629  | 
(if k = 0 then ([], x)  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
630  | 
else case f x  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
631  | 
of NONE => ([], x)  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
632  | 
| SOME (v, y) => let  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
633  | 
val k' = k - 1;  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
634  | 
val (vs, z) = anamorph f k' y  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
635  | 
in (v :: vs, z) end);  | 
| 
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
636  | 
|
| 36513 | 637  | 
datatype pred = datatype Predicate.pred  | 
638  | 
datatype seq = datatype Predicate.seq  | 
|
639  | 
||
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
640  | 
fun map f = @{code Predicate.map} f;
 | 
| 
30959
 
458e55fd0a33
fixed compilation of predicate types in ML environment
 
haftmann 
parents: 
30948 
diff
changeset
 | 
641  | 
|
| 36513 | 642  | 
fun yield (Seq f) = next (f ())  | 
643  | 
and next Empty = NONE  | 
|
644  | 
| next (Insert (x, P)) = SOME (x, P)  | 
|
645  | 
| next (Join (P, xq)) = (case yield P  | 
|
| 
30959
 
458e55fd0a33
fixed compilation of predicate types in ML environment
 
haftmann 
parents: 
30948 
diff
changeset
 | 
646  | 
of NONE => next xq  | 
| 36513 | 647  | 
| 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
 | 
648  | 
|
| 
51126
 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 
haftmann 
parents: 
51112 
diff
changeset
 | 
649  | 
fun yieldn k = anamorph yield k;  | 
| 30948 | 650  | 
|
651  | 
end;  | 
|
| 60758 | 652  | 
\<close>  | 
| 30948 | 653  | 
|
| 60758 | 654  | 
text \<open>Conversion from and to sets\<close>  | 
| 
46038
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
655  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
656  | 
definition pred_of_set :: "'a set \<Rightarrow> 'a pred" where  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
657  | 
"pred_of_set = Pred \<circ> (\<lambda>A x. x \<in> A)"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
658  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
659  | 
lemma eval_pred_of_set [simp]:  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
660  | 
"eval (pred_of_set A) x \<longleftrightarrow> x \<in>A"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
661  | 
by (simp add: pred_of_set_def)  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
662  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
663  | 
definition set_of_pred :: "'a pred \<Rightarrow> 'a set" where  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
664  | 
"set_of_pred = Collect \<circ> eval"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
665  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
666  | 
lemma member_set_of_pred [simp]:  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
667  | 
"x \<in> set_of_pred P \<longleftrightarrow> Predicate.eval P x"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
668  | 
by (simp add: set_of_pred_def)  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
669  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
670  | 
definition set_of_seq :: "'a seq \<Rightarrow> 'a set" where  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
671  | 
"set_of_seq = set_of_pred \<circ> pred_of_seq"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
672  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
673  | 
lemma member_set_of_seq [simp]:  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
674  | 
"x \<in> set_of_seq xq = Predicate.member xq x"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
675  | 
by (simp add: set_of_seq_def eval_member)  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
676  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
677  | 
lemma of_pred_code [code]:  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
678  | 
"set_of_pred (Predicate.Seq f) = (case f () of  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
679  | 
     Predicate.Empty \<Rightarrow> {}
 | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
680  | 
| Predicate.Insert x P \<Rightarrow> insert x (set_of_pred P)  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
681  | 
| Predicate.Join P xq \<Rightarrow> set_of_pred P \<union> set_of_seq xq)"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
682  | 
by (auto split: seq.split simp add: eval_code)  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
683  | 
|
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
684  | 
lemma of_seq_code [code]:  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
685  | 
  "set_of_seq Predicate.Empty = {}"
 | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
686  | 
"set_of_seq (Predicate.Insert x P) = insert x (set_of_pred P)"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
687  | 
"set_of_seq (Predicate.Join P xq) = set_of_pred P \<union> set_of_seq xq"  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
688  | 
by auto  | 
| 
 
bb2f7488a0f1
conversions from sets to predicates and vice versa; extensionality on predicates
 
haftmann 
parents: 
45970 
diff
changeset
 | 
689  | 
|
| 60758 | 690  | 
text \<open>Lazy Evaluation of an indexed function\<close>  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
691  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
692  | 
function iterate_upto :: "(natural \<Rightarrow> 'a) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> 'a Predicate.pred"  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
693  | 
where  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
694  | 
"iterate_upto f n m =  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
695  | 
Predicate.Seq (%u. if n > m then Predicate.Empty  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
696  | 
else Predicate.Insert (f n) (iterate_upto f (n + 1) m))"  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
697  | 
by pat_completeness auto  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
698  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
699  | 
termination by (relation "measure (%(f, n, m). nat_of_natural (m + 1 - n))")  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
700  | 
(auto simp add: less_natural_def)  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
701  | 
|
| 60758 | 702  | 
text \<open>Misc\<close>  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
703  | 
|
| 47399 | 704  | 
declare Inf_set_fold [where 'a = "'a Predicate.pred", code]  | 
705  | 
declare Sup_set_fold [where 'a = "'a Predicate.pred", code]  | 
|
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
706  | 
|
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
707  | 
(* FIXME: better implement conversion by bisection *)  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
708  | 
|
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
709  | 
lemma pred_of_set_fold_sup:  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
710  | 
assumes "finite A"  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
711  | 
shows "pred_of_set A = Finite_Set.fold sup bot (Predicate.single ` A)" (is "?lhs = ?rhs")  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
712  | 
proof (rule sym)  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
713  | 
interpret comp_fun_idem "sup :: 'a Predicate.pred \<Rightarrow> 'a Predicate.pred \<Rightarrow> 'a Predicate.pred"  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
714  | 
by (fact comp_fun_idem_sup)  | 
| 60758 | 715  | 
from \<open>finite A\<close> show "?rhs = ?lhs" by (induct A) (auto intro!: pred_eqI)  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
716  | 
qed  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
717  | 
|
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
718  | 
lemma pred_of_set_set_fold_sup:  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
719  | 
"pred_of_set (set xs) = fold sup (List.map Predicate.single xs) bot"  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
720  | 
proof -  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
721  | 
interpret comp_fun_idem "sup :: 'a Predicate.pred \<Rightarrow> 'a Predicate.pred \<Rightarrow> 'a Predicate.pred"  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
722  | 
by (fact comp_fun_idem_sup)  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
723  | 
show ?thesis by (simp add: pred_of_set_fold_sup fold_set_fold [symmetric])  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
724  | 
qed  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
725  | 
|
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
726  | 
lemma pred_of_set_set_foldr_sup [code]:  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
727  | 
"pred_of_set (set xs) = foldr sup (List.map Predicate.single xs) bot"  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
728  | 
by (simp add: pred_of_set_set_fold_sup ac_simps foldr_fold fun_eq_iff)  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
729  | 
|
| 30328 | 730  | 
no_notation  | 
| 62026 | 731  | 
bind (infixl "\<bind>" 70)  | 
| 30328 | 732  | 
|
| 
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
 | 
733  | 
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
 | 
734  | 
hide_const (open) Pred eval single bind is_empty singleton if_pred not_pred holds  | 
| 53943 | 735  | 
Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map the  | 
| 
46664
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
736  | 
iterate_upto  | 
| 
 
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
 
haftmann 
parents: 
46638 
diff
changeset
 | 
737  | 
hide_fact (open) null_def member_def  | 
| 30328 | 738  | 
|
739  | 
end  |