15737
|
1 |
(* Title: HOL/Library/List_lexord.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Norbert Voelker
|
|
4 |
*)
|
|
5 |
|
17200
|
6 |
header {* Lexicographic order on lists *}
|
15737
|
7 |
|
|
8 |
theory List_lexord
|
|
9 |
imports Main
|
|
10 |
begin
|
|
11 |
|
21458
|
12 |
instance list :: (ord) ord
|
17200
|
13 |
list_le_def: "(xs::('a::ord) list) \<le> ys \<equiv> (xs < ys \<or> xs = ys)"
|
21458
|
14 |
list_less_def: "(xs::('a::ord) list) < ys \<equiv> (xs, ys) \<in> lexord {(u,v). u < v}" ..
|
15737
|
15 |
|
|
16 |
lemmas list_ord_defs = list_less_def list_le_def
|
|
17 |
|
17200
|
18 |
instance list :: (order) order
|
15737
|
19 |
apply (intro_classes, unfold list_ord_defs)
|
17200
|
20 |
apply (rule disjI2, safe)
|
|
21 |
apply (blast intro: lexord_trans transI order_less_trans)
|
|
22 |
apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
|
|
23 |
apply simp
|
|
24 |
apply (blast intro: lexord_trans transI order_less_trans)
|
15737
|
25 |
apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
|
|
26 |
apply simp
|
17200
|
27 |
apply assumption
|
|
28 |
done
|
15737
|
29 |
|
21458
|
30 |
instance list :: (linorder) linorder
|
15737
|
31 |
apply (intro_classes, unfold list_le_def list_less_def, safe)
|
17200
|
32 |
apply (cut_tac x = x and y = y and r = "{(a,b). a < b}" in lexord_linear)
|
|
33 |
apply force
|
|
34 |
apply simp
|
|
35 |
done
|
15737
|
36 |
|
22177
|
37 |
lemma not_less_Nil [simp]: "\<not> (x < [])"
|
17200
|
38 |
by (unfold list_less_def) simp
|
15737
|
39 |
|
22177
|
40 |
lemma Nil_less_Cons [simp]: "[] < a # x"
|
17200
|
41 |
by (unfold list_less_def) simp
|
15737
|
42 |
|
22177
|
43 |
lemma Cons_less_Cons [simp]: "a # x < b # y \<longleftrightarrow> a < b \<or> a = b \<and> x < y"
|
17200
|
44 |
by (unfold list_less_def) simp
|
15737
|
45 |
|
22177
|
46 |
lemma le_Nil [simp]: "x \<le> [] \<longleftrightarrow> x = []"
|
|
47 |
by (unfold list_ord_defs, cases x) auto
|
|
48 |
|
|
49 |
lemma Nil_le_Cons [simp]: "[] \<le> x"
|
17200
|
50 |
by (unfold list_ord_defs, cases x) auto
|
15737
|
51 |
|
22177
|
52 |
lemma Cons_le_Cons [simp]: "a # x \<le> b # y \<longleftrightarrow> a < b \<or> a = b \<and> x \<le> y"
|
17200
|
53 |
by (unfold list_ord_defs) auto
|
15737
|
54 |
|
22177
|
55 |
lemma less_code [code func]:
|
|
56 |
"xs < ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
|
|
57 |
"[] < (x\<Colon>'a\<Colon>{eq, order}) # xs \<longleftrightarrow> True"
|
|
58 |
"(x\<Colon>'a\<Colon>{eq, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
|
|
59 |
by simp_all
|
|
60 |
|
|
61 |
lemma less_eq_code [code func]:
|
|
62 |
"x # xs \<le> ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
|
|
63 |
"[] \<le> (xs\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> True"
|
|
64 |
"(x\<Colon>'a\<Colon>{eq, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
|
|
65 |
by simp_all
|
|
66 |
|
17200
|
67 |
end
|