| author | wenzelm | 
| Thu, 11 Apr 2024 12:05:01 +0200 | |
| changeset 80109 | dbcd6dc7f70f | 
| parent 69275 | 9bbd5497befd | 
| child 81706 | 7beb0cf38292 | 
| permissions | -rw-r--r-- | 
| 30246 | 1  | 
(* Title: HOL/Option.thy  | 
2  | 
Author: Folklore  | 
|
3  | 
*)  | 
|
4  | 
||
| 60758 | 5  | 
section \<open>Datatype option\<close>  | 
| 30246 | 6  | 
|
7  | 
theory Option  | 
|
| 66364 | 8  | 
imports Lifting  | 
| 30246 | 9  | 
begin  | 
10  | 
||
| 58310 | 11  | 
datatype 'a option =  | 
| 57091 | 12  | 
None  | 
| 55406 | 13  | 
| Some (the: 'a)  | 
| 
57123
 
b5324647e0f1
tuned whitespace, to make datatype definitions slightly less intimidating
 
blanchet 
parents: 
57091 
diff
changeset
 | 
14  | 
|
| 
55531
 
601ca8efa000
renamed 'datatype_new_compat' to 'datatype_compat'
 
blanchet 
parents: 
55518 
diff
changeset
 | 
15  | 
datatype_compat option  | 
| 
55404
 
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
 
blanchet 
parents: 
55129 
diff
changeset
 | 
16  | 
|
| 55406 | 17  | 
lemma [case_names None Some, cases type: option]:  | 
| 61799 | 18  | 
\<comment> \<open>for backward compatibility -- names of variables differ\<close>  | 
| 
55417
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
55414 
diff
changeset
 | 
19  | 
"(y = None \<Longrightarrow> P) \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> P) \<Longrightarrow> P"  | 
| 61066 | 20  | 
by (rule option.exhaust)  | 
| 55406 | 21  | 
|
22  | 
lemma [case_names None Some, induct type: option]:  | 
|
| 61799 | 23  | 
\<comment> \<open>for backward compatibility -- names of variables differ\<close>  | 
| 55406 | 24  | 
"P None \<Longrightarrow> (\<And>option. P (Some option)) \<Longrightarrow> P option"  | 
| 61066 | 25  | 
by (rule option.induct)  | 
| 55406 | 26  | 
|
| 60758 | 27  | 
text \<open>Compatibility:\<close>  | 
28  | 
setup \<open>Sign.mandatory_path "option"\<close>  | 
|
| 
55404
 
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
 
blanchet 
parents: 
55129 
diff
changeset
 | 
29  | 
lemmas inducts = option.induct  | 
| 
 
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
 
blanchet 
parents: 
55129 
diff
changeset
 | 
30  | 
lemmas cases = option.case  | 
| 60758 | 31  | 
setup \<open>Sign.parent_path\<close>  | 
| 30246 | 32  | 
|
| 61066 | 33  | 
lemma not_None_eq [iff]: "x \<noteq> None \<longleftrightarrow> (\<exists>y. x = Some y)"  | 
| 30246 | 34  | 
by (induct x) auto  | 
35  | 
||
| 61066 | 36  | 
lemma not_Some_eq [iff]: "(\<forall>y. x \<noteq> Some y) \<longleftrightarrow> x = None"  | 
| 30246 | 37  | 
by (induct x) auto  | 
38  | 
||
| 68460 | 39  | 
lemma comp_the_Some[simp]: "the o Some = id"  | 
40  | 
by auto  | 
|
41  | 
||
| 61066 | 42  | 
text \<open>Although it may appear that both of these equalities are helpful  | 
| 30246 | 43  | 
only when applied to assumptions, in practice it seems better to give  | 
| 60758 | 44  | 
them the uniform iff attribute.\<close>  | 
| 30246 | 45  | 
|
| 31080 | 46  | 
lemma inj_Some [simp]: "inj_on Some A"  | 
| 61066 | 47  | 
by (rule inj_onI) simp  | 
| 31080 | 48  | 
|
| 
55404
 
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
 
blanchet 
parents: 
55129 
diff
changeset
 | 
49  | 
lemma case_optionE:  | 
| 61066 | 50  | 
assumes c: "(case x of None \<Rightarrow> P | Some y \<Rightarrow> Q y)"  | 
| 30246 | 51  | 
obtains  | 
52  | 
(None) "x = None" and P  | 
|
53  | 
| (Some) y where "x = Some y" and "Q y"  | 
|
54  | 
using c by (cases x) simp_all  | 
|
55  | 
||
| 53010 | 56  | 
lemma split_option_all: "(\<forall>x. P x) \<longleftrightarrow> P None \<and> (\<forall>x. P (Some x))"  | 
| 61066 | 57  | 
by (auto intro: option.induct)  | 
| 53010 | 58  | 
|
59  | 
lemma split_option_ex: "(\<exists>x. P x) \<longleftrightarrow> P None \<or> (\<exists>x. P (Some x))"  | 
|
| 61066 | 60  | 
using split_option_all[of "\<lambda>x. \<not> P x"] by blast  | 
| 53010 | 61  | 
|
| 31080 | 62  | 
lemma UNIV_option_conv: "UNIV = insert None (range Some)"  | 
| 61066 | 63  | 
by (auto intro: classical)  | 
| 31080 | 64  | 
|
| 59522 | 65  | 
lemma rel_option_None1 [simp]: "rel_option P None x \<longleftrightarrow> x = None"  | 
| 61066 | 66  | 
by (cases x) simp_all  | 
| 59522 | 67  | 
|
68  | 
lemma rel_option_None2 [simp]: "rel_option P x None \<longleftrightarrow> x = None"  | 
|
| 61066 | 69  | 
by (cases x) simp_all  | 
| 59522 | 70  | 
|
| 61630 | 71  | 
lemma option_rel_Some1: "rel_option A (Some x) y \<longleftrightarrow> (\<exists>y'. y = Some y' \<and> A x y')" (* Option *)  | 
72  | 
by(cases y) simp_all  | 
|
73  | 
||
74  | 
lemma option_rel_Some2: "rel_option A x (Some y) \<longleftrightarrow> (\<exists>x'. x = Some x' \<and> A x' y)" (* Option *)  | 
|
75  | 
by(cases x) simp_all  | 
|
76  | 
||
| 61066 | 77  | 
lemma rel_option_inf: "inf (rel_option A) (rel_option B) = rel_option (inf A B)"  | 
78  | 
(is "?lhs = ?rhs")  | 
|
79  | 
proof (rule antisym)  | 
|
80  | 
show "?lhs \<le> ?rhs" by (auto elim: option.rel_cases)  | 
|
81  | 
show "?rhs \<le> ?lhs" by (auto elim: option.rel_mono_strong)  | 
|
82  | 
qed  | 
|
| 59522 | 83  | 
|
84  | 
lemma rel_option_reflI:  | 
|
85  | 
"(\<And>x. x \<in> set_option y \<Longrightarrow> P x x) \<Longrightarrow> rel_option P y y"  | 
|
| 61066 | 86  | 
by (cases y) auto  | 
| 59522 | 87  | 
|
| 59523 | 88  | 
|
| 60758 | 89  | 
subsubsection \<open>Operations\<close>  | 
| 30246 | 90  | 
|
| 61066 | 91  | 
lemma ospec [dest]: "(\<forall>x\<in>set_option A. P x) \<Longrightarrow> A = Some x \<Longrightarrow> P x"  | 
| 30246 | 92  | 
by simp  | 
93  | 
||
| 60758 | 94  | 
setup \<open>map_theory_claset (fn ctxt => ctxt addSD2 ("ospec", @{thm ospec}))\<close>
 | 
| 30246 | 95  | 
|
| 61066 | 96  | 
lemma elem_set [iff]: "(x \<in> set_option xo) = (xo = Some x)"  | 
| 30246 | 97  | 
by (cases xo) auto  | 
98  | 
||
| 
55518
 
1ddb2edf5ceb
folded 'Option.set' into BNF-generated 'set_option'
 
blanchet 
parents: 
55467 
diff
changeset
 | 
99  | 
lemma set_empty_eq [simp]: "(set_option xo = {}) = (xo = None)"
 | 
| 30246 | 100  | 
by (cases xo) auto  | 
101  | 
||
| 61066 | 102  | 
lemma map_option_case: "map_option f y = (case y of None \<Rightarrow> None | Some x \<Rightarrow> Some (f x))"  | 
| 55466 | 103  | 
by (auto split: option.split)  | 
| 30246 | 104  | 
|
| 61066 | 105  | 
lemma map_option_is_None [iff]: "(map_option f opt = None) = (opt = None)"  | 
| 63648 | 106  | 
by (simp add: map_option_case split: option.split)  | 
| 30246 | 107  | 
|
| 61630 | 108  | 
lemma None_eq_map_option_iff [iff]: "None = map_option f x \<longleftrightarrow> x = None"  | 
109  | 
by(cases x) simp_all  | 
|
110  | 
||
| 61066 | 111  | 
lemma map_option_eq_Some [iff]: "(map_option f xo = Some y) = (\<exists>z. xo = Some z \<and> f z = y)"  | 
| 63648 | 112  | 
by (simp add: map_option_case split: option.split)  | 
| 30246 | 113  | 
|
| 55466 | 114  | 
lemma map_option_o_case_sum [simp]:  | 
| 67091 | 115  | 
"map_option f \<circ> case_sum g h = case_sum (map_option f \<circ> g) (map_option f \<circ> h)"  | 
| 55466 | 116  | 
by (rule o_case_sum)  | 
| 30246 | 117  | 
|
| 55466 | 118  | 
lemma map_option_cong: "x = y \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> f a = g a) \<Longrightarrow> map_option f x = map_option g y"  | 
| 61066 | 119  | 
by (cases x) auto  | 
| 46526 | 120  | 
|
| 61630 | 121  | 
lemma map_option_idI: "(\<And>y. y \<in> set_option x \<Longrightarrow> f y = y) \<Longrightarrow> map_option f x = x"  | 
122  | 
by(cases x)(simp_all)  | 
|
123  | 
||
| 59521 | 124  | 
functor map_option: map_option  | 
| 61066 | 125  | 
by (simp_all add: option.map_comp fun_eq_iff option.map_id)  | 
| 40609 | 126  | 
|
| 61066 | 127  | 
lemma case_map_option [simp]: "case_option g h (map_option f x) = case_option g (h \<circ> f) x"  | 
| 51096 | 128  | 
by (cases x) simp_all  | 
129  | 
||
| 61630 | 130  | 
lemma None_notin_image_Some [simp]: "None \<notin> Some ` A"  | 
131  | 
by auto  | 
|
132  | 
||
133  | 
lemma notin_range_Some: "x \<notin> range Some \<longleftrightarrow> x = None"  | 
|
134  | 
by(cases x) auto  | 
|
135  | 
||
| 58916 | 136  | 
lemma rel_option_iff:  | 
137  | 
"rel_option R x y = (case (x, y) of (None, None) \<Rightarrow> True  | 
|
138  | 
| (Some x, Some y) \<Rightarrow> R x y  | 
|
139  | 
| _ \<Rightarrow> False)"  | 
|
| 61066 | 140  | 
by (auto split: prod.split option.split)  | 
| 58916 | 141  | 
|
| 63194 | 142  | 
|
143  | 
definition combine_options :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a option \<Rightarrow> 'a option \<Rightarrow> 'a option"
 | 
|
144  | 
where "combine_options f x y =  | 
|
145  | 
(case x of None \<Rightarrow> y | Some x \<Rightarrow> (case y of None \<Rightarrow> Some x | Some y \<Rightarrow> Some (f x y)))"  | 
|
146  | 
||
147  | 
lemma combine_options_simps [simp]:  | 
|
148  | 
"combine_options f None y = y"  | 
|
149  | 
"combine_options f x None = x"  | 
|
150  | 
"combine_options f (Some a) (Some b) = Some (f a b)"  | 
|
151  | 
by (simp_all add: combine_options_def split: option.splits)  | 
|
152  | 
||
153  | 
lemma combine_options_cases [case_names None1 None2 Some]:  | 
|
154  | 
"(x = None \<Longrightarrow> P x y) \<Longrightarrow> (y = None \<Longrightarrow> P x y) \<Longrightarrow>  | 
|
155  | 
(\<And>a b. x = Some a \<Longrightarrow> y = Some b \<Longrightarrow> P x y) \<Longrightarrow> P x y"  | 
|
156  | 
by (cases x; cases y) simp_all  | 
|
157  | 
||
158  | 
lemma combine_options_commute:  | 
|
159  | 
"(\<And>x y. f x y = f y x) \<Longrightarrow> combine_options f x y = combine_options f y x"  | 
|
160  | 
using combine_options_cases[of x ]  | 
|
161  | 
by (induction x y rule: combine_options_cases) simp_all  | 
|
162  | 
||
163  | 
lemma combine_options_assoc:  | 
|
164  | 
"(\<And>x y z. f (f x y) z = f x (f y z)) \<Longrightarrow>  | 
|
165  | 
combine_options f (combine_options f x y) z =  | 
|
166  | 
combine_options f x (combine_options f y z)"  | 
|
167  | 
by (auto simp: combine_options_def split: option.splits)  | 
|
168  | 
||
169  | 
lemma combine_options_left_commute:  | 
|
170  | 
"(\<And>x y. f x y = f y x) \<Longrightarrow> (\<And>x y z. f (f x y) z = f x (f y z)) \<Longrightarrow>  | 
|
171  | 
combine_options f y (combine_options f x z) =  | 
|
172  | 
combine_options f x (combine_options f y z)"  | 
|
173  | 
by (auto simp: combine_options_def split: option.splits)  | 
|
174  | 
||
175  | 
lemmas combine_options_ac =  | 
|
176  | 
combine_options_commute combine_options_assoc combine_options_left_commute  | 
|
177  | 
||
178  | 
||
| 61068 | 179  | 
context  | 
180  | 
begin  | 
|
181  | 
||
182  | 
qualified definition is_none :: "'a option \<Rightarrow> bool"  | 
|
| 61066 | 183  | 
where [code_post]: "is_none x \<longleftrightarrow> x = None"  | 
| 59522 | 184  | 
|
185  | 
lemma is_none_simps [simp]:  | 
|
186  | 
"is_none None"  | 
|
187  | 
"\<not> is_none (Some x)"  | 
|
| 61066 | 188  | 
by (simp_all add: is_none_def)  | 
| 59522 | 189  | 
|
190  | 
lemma is_none_code [code]:  | 
|
191  | 
"is_none None = True"  | 
|
192  | 
"is_none (Some x) = False"  | 
|
| 61066 | 193  | 
by simp_all  | 
| 59522 | 194  | 
|
195  | 
lemma rel_option_unfold:  | 
|
196  | 
"rel_option R x y \<longleftrightarrow>  | 
|
197  | 
(is_none x \<longleftrightarrow> is_none y) \<and> (\<not> is_none x \<longrightarrow> \<not> is_none y \<longrightarrow> R (the x) (the y))"  | 
|
| 61066 | 198  | 
by (simp add: rel_option_iff split: option.split)  | 
| 59522 | 199  | 
|
200  | 
lemma rel_optionI:  | 
|
201  | 
"\<lbrakk> is_none x \<longleftrightarrow> is_none y; \<lbrakk> \<not> is_none x; \<not> is_none y \<rbrakk> \<Longrightarrow> P (the x) (the y) \<rbrakk>  | 
|
202  | 
\<Longrightarrow> rel_option P x y"  | 
|
| 61066 | 203  | 
by (simp add: rel_option_unfold)  | 
| 59522 | 204  | 
|
205  | 
lemma is_none_map_option [simp]: "is_none (map_option f x) \<longleftrightarrow> is_none x"  | 
|
| 61066 | 206  | 
by (simp add: is_none_def)  | 
| 59522 | 207  | 
|
208  | 
lemma the_map_option: "\<not> is_none x \<Longrightarrow> the (map_option f x) = f (the x)"  | 
|
| 61066 | 209  | 
by (auto simp add: is_none_def)  | 
| 59522 | 210  | 
|
211  | 
||
| 61068 | 212  | 
qualified primrec bind :: "'a option \<Rightarrow> ('a \<Rightarrow> 'b option) \<Rightarrow> 'b option"
 | 
| 61066 | 213  | 
where  | 
214  | 
bind_lzero: "bind None f = None"  | 
|
215  | 
| bind_lunit: "bind (Some x) f = f x"  | 
|
| 30246 | 216  | 
|
| 59522 | 217  | 
lemma is_none_bind: "is_none (bind f g) \<longleftrightarrow> is_none f \<or> is_none (g (the f))"  | 
| 61066 | 218  | 
by (cases f) simp_all  | 
| 59522 | 219  | 
|
| 39149 | 220  | 
lemma bind_runit[simp]: "bind x Some = x"  | 
| 61066 | 221  | 
by (cases x) auto  | 
| 39149 | 222  | 
|
223  | 
lemma bind_assoc[simp]: "bind (bind x f) g = bind x (\<lambda>y. bind (f y) g)"  | 
|
| 61066 | 224  | 
by (cases x) auto  | 
| 39149 | 225  | 
|
226  | 
lemma bind_rzero[simp]: "bind x (\<lambda>x. None) = None"  | 
|
| 61066 | 227  | 
by (cases x) auto  | 
| 39149 | 228  | 
|
| 61068 | 229  | 
qualified lemma bind_cong: "x = y \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> f a = g a) \<Longrightarrow> bind x f = bind y g"  | 
| 61066 | 230  | 
by (cases x) auto  | 
| 46526 | 231  | 
|
| 61066 | 232  | 
lemma bind_split: "P (bind m f) \<longleftrightarrow> (m = None \<longrightarrow> P None) \<and> (\<forall>v. m = Some v \<longrightarrow> P (f v))"  | 
233  | 
by (cases m) auto  | 
|
| 
58895
 
de0a4a76d7aa
Added Option.bind_split{,_asm,s}
 
lammich <lammich@in.tum.de> 
parents: 
58889 
diff
changeset
 | 
234  | 
|
| 61066 | 235  | 
lemma bind_split_asm: "P (bind m f) \<longleftrightarrow> \<not> (m = None \<and> \<not> P None \<or> (\<exists>x. m = Some x \<and> \<not> P (f x)))"  | 
| 
58895
 
de0a4a76d7aa
Added Option.bind_split{,_asm,s}
 
lammich <lammich@in.tum.de> 
parents: 
58889 
diff
changeset
 | 
236  | 
by (cases m) auto  | 
| 
 
de0a4a76d7aa
Added Option.bind_split{,_asm,s}
 
lammich <lammich@in.tum.de> 
parents: 
58889 
diff
changeset
 | 
237  | 
|
| 
 
de0a4a76d7aa
Added Option.bind_split{,_asm,s}
 
lammich <lammich@in.tum.de> 
parents: 
58889 
diff
changeset
 | 
238  | 
lemmas bind_splits = bind_split bind_split_asm  | 
| 
 
de0a4a76d7aa
Added Option.bind_split{,_asm,s}
 
lammich <lammich@in.tum.de> 
parents: 
58889 
diff
changeset
 | 
239  | 
|
| 59522 | 240  | 
lemma bind_eq_Some_conv: "bind f g = Some x \<longleftrightarrow> (\<exists>y. f = Some y \<and> g y = Some x)"  | 
| 61066 | 241  | 
by (cases f) simp_all  | 
| 59522 | 242  | 
|
| 61630 | 243  | 
lemma bind_eq_None_conv: "Option.bind a f = None \<longleftrightarrow> a = None \<or> f (the a) = None"  | 
244  | 
by(cases a) simp_all  | 
|
245  | 
||
| 59522 | 246  | 
lemma map_option_bind: "map_option f (bind x g) = bind x (map_option f \<circ> g)"  | 
| 61066 | 247  | 
by (cases x) simp_all  | 
| 59522 | 248  | 
|
249  | 
lemma bind_option_cong:  | 
|
250  | 
"\<lbrakk> x = y; \<And>z. z \<in> set_option y \<Longrightarrow> f z = g z \<rbrakk> \<Longrightarrow> bind x f = bind y g"  | 
|
| 61066 | 251  | 
by (cases y) simp_all  | 
| 59522 | 252  | 
|
253  | 
lemma bind_option_cong_simp:  | 
|
254  | 
"\<lbrakk> x = y; \<And>z. z \<in> set_option y =simp=> f z = g z \<rbrakk> \<Longrightarrow> bind x f = bind y g"  | 
|
| 61066 | 255  | 
unfolding simp_implies_def by (rule bind_option_cong)  | 
| 59522 | 256  | 
|
| 61066 | 257  | 
lemma bind_option_cong_code: "x = y \<Longrightarrow> bind x f = bind y f"  | 
258  | 
by simp  | 
|
| 61068 | 259  | 
|
| 61630 | 260  | 
lemma bind_map_option: "bind (map_option f x) g = bind x (g \<circ> f)"  | 
261  | 
by(cases x) simp_all  | 
|
262  | 
||
| 69275 | 263  | 
lemma set_bind_option [simp]: "set_option (bind x f) = (\<Union>((set_option \<circ> f) ` set_option x))"  | 
| 61630 | 264  | 
by(cases x) auto  | 
265  | 
||
266  | 
lemma map_conv_bind_option: "map_option f x = Option.bind x (Some \<circ> f)"  | 
|
267  | 
by(cases x) simp_all  | 
|
268  | 
||
| 61068 | 269  | 
end  | 
270  | 
||
| 59522 | 271  | 
setup \<open>Code_Simp.map_ss (Simplifier.add_cong @{thm bind_option_cong_code})\<close>
 | 
272  | 
||
273  | 
||
| 61068 | 274  | 
context  | 
275  | 
begin  | 
|
276  | 
||
277  | 
qualified definition these :: "'a option set \<Rightarrow> 'a set"  | 
|
| 61066 | 278  | 
  where "these A = the ` {x \<in> A. x \<noteq> None}"
 | 
| 49189 | 279  | 
|
| 61066 | 280  | 
lemma these_empty [simp]: "these {} = {}"
 | 
| 49189 | 281  | 
by (simp add: these_def)  | 
282  | 
||
| 61066 | 283  | 
lemma these_insert_None [simp]: "these (insert None A) = these A"  | 
| 49189 | 284  | 
by (auto simp add: these_def)  | 
285  | 
||
| 61066 | 286  | 
lemma these_insert_Some [simp]: "these (insert (Some x) A) = insert x (these A)"  | 
| 49189 | 287  | 
proof -  | 
288  | 
  have "{y \<in> insert (Some x) A. y \<noteq> None} = insert (Some x) {y \<in> A. y \<noteq> None}"
 | 
|
289  | 
by auto  | 
|
290  | 
then show ?thesis by (simp add: these_def)  | 
|
291  | 
qed  | 
|
292  | 
||
| 61066 | 293  | 
lemma in_these_eq: "x \<in> these A \<longleftrightarrow> Some x \<in> A"  | 
| 49189 | 294  | 
proof  | 
295  | 
assume "Some x \<in> A"  | 
|
296  | 
then obtain B where "A = insert (Some x) B" by auto  | 
|
297  | 
then show "x \<in> these A" by (auto simp add: these_def intro!: image_eqI)  | 
|
298  | 
next  | 
|
299  | 
assume "x \<in> these A"  | 
|
300  | 
then show "Some x \<in> A" by (auto simp add: these_def)  | 
|
301  | 
qed  | 
|
302  | 
||
| 61066 | 303  | 
lemma these_image_Some_eq [simp]: "these (Some ` A) = A"  | 
| 49189 | 304  | 
by (auto simp add: these_def intro!: image_eqI)  | 
305  | 
||
| 61066 | 306  | 
lemma Some_image_these_eq: "Some ` these A = {x\<in>A. x \<noteq> None}"
 | 
| 49189 | 307  | 
by (auto simp add: these_def image_image intro!: image_eqI)  | 
308  | 
||
| 61066 | 309  | 
lemma these_empty_eq: "these B = {} \<longleftrightarrow> B = {} \<or> B = {None}"
 | 
| 49189 | 310  | 
by (auto simp add: these_def)  | 
311  | 
||
| 61066 | 312  | 
lemma these_not_empty_eq: "these B \<noteq> {} \<longleftrightarrow> B \<noteq> {} \<and> B \<noteq> {None}"
 | 
| 49189 | 313  | 
by (auto simp add: these_empty_eq)  | 
314  | 
||
| 61068 | 315  | 
end  | 
| 30246 | 316  | 
|
| 68011 | 317  | 
lemma finite_range_Some: "finite (range (Some :: 'a \<Rightarrow> 'a option)) = finite (UNIV :: 'a set)"  | 
318  | 
by (auto dest: finite_imageD intro: inj_Some)  | 
|
319  | 
||
| 49189 | 320  | 
|
| 60758 | 321  | 
subsection \<open>Transfer rules for the Transfer package\<close>  | 
| 58916 | 322  | 
|
| 63343 | 323  | 
context includes lifting_syntax  | 
| 58916 | 324  | 
begin  | 
| 61066 | 325  | 
|
| 58916 | 326  | 
lemma option_bind_transfer [transfer_rule]:  | 
327  | 
"(rel_option A ===> (A ===> rel_option B) ===> rel_option B)  | 
|
328  | 
Option.bind Option.bind"  | 
|
329  | 
unfolding rel_fun_def split_option_all by simp  | 
|
330  | 
||
| 59523 | 331  | 
lemma pred_option_parametric [transfer_rule]:  | 
| 67399 | 332  | 
"((A ===> (=)) ===> rel_option A ===> (=)) pred_option pred_option"  | 
| 61068 | 333  | 
by (rule rel_funI)+ (auto simp add: rel_option_unfold Option.is_none_def dest: rel_funD)  | 
| 59523 | 334  | 
|
| 58916 | 335  | 
end  | 
336  | 
||
337  | 
||
| 60758 | 338  | 
subsubsection \<open>Interaction with finite sets\<close>  | 
| 
55089
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
339  | 
|
| 
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
340  | 
lemma finite_option_UNIV [simp]:  | 
| 
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
341  | 
"finite (UNIV :: 'a option set) = finite (UNIV :: 'a set)"  | 
| 
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
342  | 
by (auto simp add: UNIV_option_conv elim: finite_imageD intro: inj_Some)  | 
| 
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
343  | 
|
| 
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
344  | 
instance option :: (finite) finite  | 
| 61066 | 345  | 
by standard (simp add: UNIV_option_conv)  | 
| 
55089
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
346  | 
|
| 
 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 
blanchet 
parents: 
53940 
diff
changeset
 | 
347  | 
|
| 60758 | 348  | 
subsubsection \<open>Code generator setup\<close>  | 
| 30246 | 349  | 
|
| 59522 | 350  | 
lemma equal_None_code_unfold [code_unfold]:  | 
| 61068 | 351  | 
"HOL.equal x None \<longleftrightarrow> Option.is_none x"  | 
352  | 
"HOL.equal None = Option.is_none"  | 
|
353  | 
by (auto simp add: equal Option.is_none_def)  | 
|
| 30246 | 354  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
355  | 
code_printing  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
356  | 
type_constructor option \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
357  | 
(SML) "_ option"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
358  | 
and (OCaml) "_ option"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
359  | 
and (Haskell) "Maybe _"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
360  | 
and (Scala) "!Option[(_)]"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
361  | 
| constant None \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
362  | 
(SML) "NONE"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
363  | 
and (OCaml) "None"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
364  | 
and (Haskell) "Nothing"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
365  | 
and (Scala) "!None"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
366  | 
| constant Some \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
367  | 
(SML) "SOME"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
368  | 
and (OCaml) "Some _"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
369  | 
and (Haskell) "Just"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
370  | 
and (Scala) "Some"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
371  | 
| class_instance option :: equal \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
372  | 
(Haskell) -  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
373  | 
| constant "HOL.equal :: 'a option \<Rightarrow> 'a option \<Rightarrow> bool" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51703 
diff
changeset
 | 
374  | 
(Haskell) infix 4 "=="  | 
| 30246 | 375  | 
|
376  | 
code_reserved SML  | 
|
377  | 
option NONE SOME  | 
|
378  | 
||
379  | 
code_reserved OCaml  | 
|
380  | 
option None Some  | 
|
381  | 
||
| 34886 | 382  | 
code_reserved Scala  | 
383  | 
Option None Some  | 
|
384  | 
||
| 30246 | 385  | 
end  |