26241
|
1 |
(* Title: HOL/Library/Option_ord.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Florian Haftmann, TU Muenchen
|
|
4 |
*)
|
|
5 |
|
26263
|
6 |
header {* Canonical order on option type *}
|
26241
|
7 |
|
|
8 |
theory Option_ord
|
27368
|
9 |
imports Plain
|
26241
|
10 |
begin
|
|
11 |
|
|
12 |
instantiation option :: (order) order
|
|
13 |
begin
|
|
14 |
|
|
15 |
definition less_eq_option where
|
28562
|
16 |
[code del]: "x \<le> y \<longleftrightarrow> (case x of None \<Rightarrow> True | Some x \<Rightarrow> (case y of None \<Rightarrow> False | Some y \<Rightarrow> x \<le> y))"
|
26241
|
17 |
|
|
18 |
definition less_option where
|
28562
|
19 |
[code del]: "x < y \<longleftrightarrow> (case y of None \<Rightarrow> False | Some y \<Rightarrow> (case x of None \<Rightarrow> True | Some x \<Rightarrow> x < y))"
|
26241
|
20 |
|
26258
|
21 |
lemma less_eq_option_None [simp]: "None \<le> x"
|
26241
|
22 |
by (simp add: less_eq_option_def)
|
|
23 |
|
26258
|
24 |
lemma less_eq_option_None_code [code]: "None \<le> x \<longleftrightarrow> True"
|
26241
|
25 |
by simp
|
|
26 |
|
26258
|
27 |
lemma less_eq_option_None_is_None: "x \<le> None \<Longrightarrow> x = None"
|
26241
|
28 |
by (cases x) (simp_all add: less_eq_option_def)
|
|
29 |
|
26258
|
30 |
lemma less_eq_option_Some_None [simp, code]: "Some x \<le> None \<longleftrightarrow> False"
|
26241
|
31 |
by (simp add: less_eq_option_def)
|
|
32 |
|
26258
|
33 |
lemma less_eq_option_Some [simp, code]: "Some x \<le> Some y \<longleftrightarrow> x \<le> y"
|
26241
|
34 |
by (simp add: less_eq_option_def)
|
|
35 |
|
26258
|
36 |
lemma less_option_None [simp, code]: "x < None \<longleftrightarrow> False"
|
26241
|
37 |
by (simp add: less_option_def)
|
|
38 |
|
26258
|
39 |
lemma less_option_None_is_Some: "None < x \<Longrightarrow> \<exists>z. x = Some z"
|
26241
|
40 |
by (cases x) (simp_all add: less_option_def)
|
|
41 |
|
26258
|
42 |
lemma less_option_None_Some [simp]: "None < Some x"
|
26241
|
43 |
by (simp add: less_option_def)
|
|
44 |
|
26258
|
45 |
lemma less_option_None_Some_code [code]: "None < Some x \<longleftrightarrow> True"
|
26241
|
46 |
by simp
|
|
47 |
|
26258
|
48 |
lemma less_option_Some [simp, code]: "Some x < Some y \<longleftrightarrow> x < y"
|
26241
|
49 |
by (simp add: less_option_def)
|
|
50 |
|
|
51 |
instance by default
|
|
52 |
(auto simp add: less_eq_option_def less_option_def split: option.splits)
|
|
53 |
|
|
54 |
end
|
|
55 |
|
|
56 |
instance option :: (linorder) linorder
|
|
57 |
by default (auto simp add: less_eq_option_def less_option_def split: option.splits)
|
|
58 |
|
|
59 |
end
|