| 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)
 | 
| 22316 |     20 |   apply safe
 | 
| 15737 |     21 |   apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
 | 
|  |     22 |   apply simp
 | 
| 17200 |     23 |   apply assumption
 | 
| 22316 |     24 |   apply (blast intro: lexord_trans transI order_less_trans)
 | 
|  |     25 |   apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
 | 
|  |     26 |   apply simp
 | 
|  |     27 |   apply (blast intro: lexord_trans transI order_less_trans)
 | 
| 17200 |     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 | 
 | 
| 22483 |     37 | instance list :: (linorder) distrib_lattice
 | 
|  |     38 |   "inf \<equiv> min"
 | 
|  |     39 |   "sup \<equiv> max"
 | 
|  |     40 |   by intro_classes
 | 
|  |     41 |     (auto simp add: inf_list_def sup_list_def min_max.sup_inf_distrib1)
 | 
|  |     42 | 
 | 
| 22177 |     43 | lemma not_less_Nil [simp]: "\<not> (x < [])"
 | 
| 17200 |     44 |   by (unfold list_less_def) simp
 | 
| 15737 |     45 | 
 | 
| 22177 |     46 | lemma Nil_less_Cons [simp]: "[] < a # x"
 | 
| 17200 |     47 |   by (unfold list_less_def) simp
 | 
| 15737 |     48 | 
 | 
| 22177 |     49 | lemma Cons_less_Cons [simp]: "a # x < b # y \<longleftrightarrow> a < b \<or> a = b \<and> x < y"
 | 
| 17200 |     50 |   by (unfold list_less_def) simp
 | 
| 15737 |     51 | 
 | 
| 22177 |     52 | lemma le_Nil [simp]: "x \<le> [] \<longleftrightarrow> x = []"
 | 
|  |     53 |   by (unfold list_ord_defs, cases x) auto
 | 
|  |     54 | 
 | 
|  |     55 | lemma Nil_le_Cons [simp]: "[] \<le> x"
 | 
| 17200 |     56 |   by (unfold list_ord_defs, cases x) auto
 | 
| 15737 |     57 | 
 | 
| 22177 |     58 | lemma Cons_le_Cons [simp]: "a # x \<le> b # y \<longleftrightarrow> a < b \<or> a = b \<and> x \<le> y"
 | 
| 17200 |     59 |   by (unfold list_ord_defs) auto
 | 
| 15737 |     60 | 
 | 
| 22177 |     61 | lemma less_code [code func]:
 | 
|  |     62 |   "xs < ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
 | 
|  |     63 |   "[] < (x\<Colon>'a\<Colon>{eq, order}) # xs \<longleftrightarrow> True"
 | 
|  |     64 |   "(x\<Colon>'a\<Colon>{eq, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
 | 
|  |     65 |   by simp_all
 | 
|  |     66 | 
 | 
|  |     67 | lemma less_eq_code [code func]:
 | 
|  |     68 |   "x # xs \<le> ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
 | 
|  |     69 |   "[] \<le> (xs\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> True"
 | 
|  |     70 |   "(x\<Colon>'a\<Colon>{eq, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
 | 
|  |     71 |   by simp_all
 | 
|  |     72 | 
 | 
| 17200 |     73 | end
 |