author | haftmann |
Fri, 05 Feb 2010 14:33:50 +0100 | |
changeset 35028 | 108662d50512 |
parent 30663 | 0b6aff7451b2 |
child 37474 | ce943f9edf5e |
permissions | -rw-r--r-- |
15737 | 1 |
(* Title: HOL/Library/List_lexord.thy |
2 |
Author: Norbert Voelker |
|
3 |
*) |
|
4 |
||
17200 | 5 |
header {* Lexicographic order on lists *} |
15737 | 6 |
|
7 |
theory List_lexord |
|
30663
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
haftmann
parents:
28562
diff
changeset
|
8 |
imports List Main |
15737 | 9 |
begin |
10 |
||
25764 | 11 |
instantiation list :: (ord) ord |
12 |
begin |
|
13 |
||
14 |
definition |
|
28562 | 15 |
list_less_def [code del]: "(xs::('a::ord) list) < ys \<longleftrightarrow> (xs, ys) \<in> lexord {(u,v). u < v}" |
25764 | 16 |
|
17 |
definition |
|
28562 | 18 |
list_le_def [code del]: "(xs::('a::ord) list) \<le> ys \<longleftrightarrow> (xs < ys \<or> xs = ys)" |
25764 | 19 |
|
20 |
instance .. |
|
21 |
||
22 |
end |
|
15737 | 23 |
|
17200 | 24 |
instance list :: (order) order |
27682 | 25 |
proof |
26 |
fix xs :: "'a list" |
|
27 |
show "xs \<le> xs" by (simp add: list_le_def) |
|
28 |
next |
|
29 |
fix xs ys zs :: "'a list" |
|
30 |
assume "xs \<le> ys" and "ys \<le> zs" |
|
31 |
then show "xs \<le> zs" by (auto simp add: list_le_def list_less_def) |
|
32 |
(rule lexord_trans, auto intro: transI) |
|
33 |
next |
|
34 |
fix xs ys :: "'a list" |
|
35 |
assume "xs \<le> ys" and "ys \<le> xs" |
|
36 |
then show "xs = ys" apply (auto simp add: list_le_def list_less_def) |
|
37 |
apply (rule lexord_irreflexive [THEN notE]) |
|
38 |
defer |
|
39 |
apply (rule lexord_trans) apply (auto intro: transI) done |
|
40 |
next |
|
41 |
fix xs ys :: "'a list" |
|
42 |
show "xs < ys \<longleftrightarrow> xs \<le> ys \<and> \<not> ys \<le> xs" |
|
43 |
apply (auto simp add: list_less_def list_le_def) |
|
44 |
defer |
|
45 |
apply (rule lexord_irreflexive [THEN notE]) |
|
46 |
apply auto |
|
47 |
apply (rule lexord_irreflexive [THEN notE]) |
|
48 |
defer |
|
49 |
apply (rule lexord_trans) apply (auto intro: transI) done |
|
50 |
qed |
|
15737 | 51 |
|
21458 | 52 |
instance list :: (linorder) linorder |
27682 | 53 |
proof |
54 |
fix xs ys :: "'a list" |
|
55 |
have "(xs, ys) \<in> lexord {(u, v). u < v} \<or> xs = ys \<or> (ys, xs) \<in> lexord {(u, v). u < v}" |
|
56 |
by (rule lexord_linear) auto |
|
57 |
then show "xs \<le> ys \<or> ys \<le> xs" |
|
58 |
by (auto simp add: list_le_def list_less_def) |
|
59 |
qed |
|
15737 | 60 |
|
25764 | 61 |
instantiation list :: (linorder) distrib_lattice |
62 |
begin |
|
63 |
||
64 |
definition |
|
28562 | 65 |
[code del]: "(inf \<Colon> 'a list \<Rightarrow> _) = min" |
25764 | 66 |
|
67 |
definition |
|
28562 | 68 |
[code del]: "(sup \<Colon> 'a list \<Rightarrow> _) = max" |
25764 | 69 |
|
70 |
instance |
|
22483 | 71 |
by intro_classes |
72 |
(auto simp add: inf_list_def sup_list_def min_max.sup_inf_distrib1) |
|
73 |
||
25764 | 74 |
end |
75 |
||
22177 | 76 |
lemma not_less_Nil [simp]: "\<not> (x < [])" |
17200 | 77 |
by (unfold list_less_def) simp |
15737 | 78 |
|
22177 | 79 |
lemma Nil_less_Cons [simp]: "[] < a # x" |
17200 | 80 |
by (unfold list_less_def) simp |
15737 | 81 |
|
22177 | 82 |
lemma Cons_less_Cons [simp]: "a # x < b # y \<longleftrightarrow> a < b \<or> a = b \<and> x < y" |
17200 | 83 |
by (unfold list_less_def) simp |
15737 | 84 |
|
22177 | 85 |
lemma le_Nil [simp]: "x \<le> [] \<longleftrightarrow> x = []" |
25502 | 86 |
by (unfold list_le_def, cases x) auto |
22177 | 87 |
|
88 |
lemma Nil_le_Cons [simp]: "[] \<le> x" |
|
25502 | 89 |
by (unfold list_le_def, cases x) auto |
15737 | 90 |
|
22177 | 91 |
lemma Cons_le_Cons [simp]: "a # x \<le> b # y \<longleftrightarrow> a < b \<or> a = b \<and> x \<le> y" |
25502 | 92 |
by (unfold list_le_def) auto |
15737 | 93 |
|
28562 | 94 |
lemma less_code [code]: |
22177 | 95 |
"xs < ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False" |
96 |
"[] < (x\<Colon>'a\<Colon>{eq, order}) # xs \<longleftrightarrow> True" |
|
97 |
"(x\<Colon>'a\<Colon>{eq, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys" |
|
98 |
by simp_all |
|
99 |
||
28562 | 100 |
lemma less_eq_code [code]: |
22177 | 101 |
"x # xs \<le> ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False" |
102 |
"[] \<le> (xs\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> True" |
|
103 |
"(x\<Colon>'a\<Colon>{eq, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys" |
|
104 |
by simp_all |
|
105 |
||
17200 | 106 |
end |