| 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
 | 
| 25595 |      9 | imports List
 | 
| 15737 |     10 | begin
 | 
|  |     11 | 
 | 
| 25764 |     12 | instantiation list :: (ord) ord
 | 
|  |     13 | begin
 | 
|  |     14 | 
 | 
|  |     15 | definition
 | 
| 25502 |     16 |   list_less_def [code func del]: "(xs::('a::ord) list) < ys \<longleftrightarrow> (xs, ys) \<in> lexord {(u,v). u < v}"
 | 
| 25764 |     17 | 
 | 
|  |     18 | definition
 | 
|  |     19 |   list_le_def [code func del]: "(xs::('a::ord) list) \<le> ys \<longleftrightarrow> (xs < ys \<or> xs = ys)"
 | 
|  |     20 | 
 | 
|  |     21 | instance ..
 | 
|  |     22 | 
 | 
|  |     23 | end
 | 
| 15737 |     24 | 
 | 
| 17200 |     25 | instance list :: (order) order
 | 
| 25502 |     26 |   apply (intro_classes, unfold list_less_def list_le_def)
 | 
| 22316 |     27 |   apply safe
 | 
| 15737 |     28 |   apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
 | 
|  |     29 |   apply simp
 | 
| 17200 |     30 |   apply assumption
 | 
| 22316 |     31 |   apply (blast intro: lexord_trans transI order_less_trans)
 | 
|  |     32 |   apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
 | 
|  |     33 |   apply simp
 | 
|  |     34 |   apply (blast intro: lexord_trans transI order_less_trans)
 | 
| 17200 |     35 |   done
 | 
| 15737 |     36 | 
 | 
| 21458 |     37 | instance list :: (linorder) linorder
 | 
| 15737 |     38 |   apply (intro_classes, unfold list_le_def list_less_def, safe)
 | 
| 17200 |     39 |   apply (cut_tac x = x and y = y and  r = "{(a,b). a < b}"  in lexord_linear)
 | 
|  |     40 |    apply force
 | 
|  |     41 |   apply simp
 | 
|  |     42 |   done
 | 
| 15737 |     43 | 
 | 
| 25764 |     44 | instantiation list :: (linorder) distrib_lattice
 | 
|  |     45 | begin
 | 
|  |     46 | 
 | 
|  |     47 | definition
 | 
| 25502 |     48 |   [code func del]: "(inf \<Colon> 'a list \<Rightarrow> _) = min"
 | 
| 25764 |     49 | 
 | 
|  |     50 | definition
 | 
| 25502 |     51 |   [code func del]: "(sup \<Colon> 'a list \<Rightarrow> _) = max"
 | 
| 25764 |     52 | 
 | 
|  |     53 | instance
 | 
| 22483 |     54 |   by intro_classes
 | 
|  |     55 |     (auto simp add: inf_list_def sup_list_def min_max.sup_inf_distrib1)
 | 
|  |     56 | 
 | 
| 25764 |     57 | end
 | 
|  |     58 | 
 | 
| 22177 |     59 | lemma not_less_Nil [simp]: "\<not> (x < [])"
 | 
| 17200 |     60 |   by (unfold list_less_def) simp
 | 
| 15737 |     61 | 
 | 
| 22177 |     62 | lemma Nil_less_Cons [simp]: "[] < a # x"
 | 
| 17200 |     63 |   by (unfold list_less_def) simp
 | 
| 15737 |     64 | 
 | 
| 22177 |     65 | lemma Cons_less_Cons [simp]: "a # x < b # y \<longleftrightarrow> a < b \<or> a = b \<and> x < y"
 | 
| 17200 |     66 |   by (unfold list_less_def) simp
 | 
| 15737 |     67 | 
 | 
| 22177 |     68 | lemma le_Nil [simp]: "x \<le> [] \<longleftrightarrow> x = []"
 | 
| 25502 |     69 |   by (unfold list_le_def, cases x) auto
 | 
| 22177 |     70 | 
 | 
|  |     71 | lemma Nil_le_Cons [simp]: "[] \<le> x"
 | 
| 25502 |     72 |   by (unfold list_le_def, cases x) auto
 | 
| 15737 |     73 | 
 | 
| 22177 |     74 | lemma Cons_le_Cons [simp]: "a # x \<le> b # y \<longleftrightarrow> a < b \<or> a = b \<and> x \<le> y"
 | 
| 25502 |     75 |   by (unfold list_le_def) auto
 | 
| 15737 |     76 | 
 | 
| 22177 |     77 | lemma less_code [code func]:
 | 
|  |     78 |   "xs < ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
 | 
|  |     79 |   "[] < (x\<Colon>'a\<Colon>{eq, order}) # xs \<longleftrightarrow> True"
 | 
|  |     80 |   "(x\<Colon>'a\<Colon>{eq, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
 | 
|  |     81 |   by simp_all
 | 
|  |     82 | 
 | 
|  |     83 | lemma less_eq_code [code func]:
 | 
|  |     84 |   "x # xs \<le> ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
 | 
|  |     85 |   "[] \<le> (xs\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> True"
 | 
|  |     86 |   "(x\<Colon>'a\<Colon>{eq, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
 | 
|  |     87 |   by simp_all
 | 
|  |     88 | 
 | 
| 17200 |     89 | end
 |