| 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 | 
 | 
|  |     12 | instance list :: (ord) ord ..
 | 
| 17200 |     13 | defs (overloaded)
 | 
|  |     14 |   list_le_def:  "(xs::('a::ord) list) \<le> ys \<equiv> (xs < ys \<or> xs = ys)"
 | 
|  |     15 |   list_less_def: "(xs::('a::ord) list) < ys \<equiv> (xs, ys) \<in> lexord {(u,v). u < v}"
 | 
| 15737 |     16 | 
 | 
|  |     17 | lemmas list_ord_defs = list_less_def list_le_def
 | 
|  |     18 | 
 | 
| 17200 |     19 | instance list :: (order) order
 | 
| 15737 |     20 |   apply (intro_classes, unfold list_ord_defs)
 | 
| 17200 |     21 |      apply (rule disjI2, safe)
 | 
|  |     22 |     apply (blast intro: lexord_trans transI order_less_trans)
 | 
|  |     23 |    apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
 | 
|  |     24 |     apply simp
 | 
|  |     25 |    apply (blast intro: lexord_trans transI order_less_trans)
 | 
| 15737 |     26 |   apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
 | 
|  |     27 |   apply simp
 | 
| 17200 |     28 |   apply assumption
 | 
|  |     29 |   done
 | 
| 15737 |     30 | 
 | 
|  |     31 | instance list::(linorder)linorder
 | 
|  |     32 |   apply (intro_classes, unfold list_le_def list_less_def, safe)
 | 
| 17200 |     33 |   apply (cut_tac x = x and y = y and  r = "{(a,b). a < b}"  in lexord_linear)
 | 
|  |     34 |    apply force
 | 
|  |     35 |   apply simp
 | 
|  |     36 |   done
 | 
| 15737 |     37 | 
 | 
| 17200 |     38 | lemma not_less_Nil[simp]: "~(x < [])"
 | 
|  |     39 |   by (unfold list_less_def) simp
 | 
| 15737 |     40 | 
 | 
| 17200 |     41 | lemma Nil_less_Cons[simp]: "[] < a # x"
 | 
|  |     42 |   by (unfold list_less_def) simp
 | 
| 15737 |     43 | 
 | 
| 17200 |     44 | lemma Cons_less_Cons[simp]: "(a # x < b # y) = (a < b | a = b & x < y)"
 | 
|  |     45 |   by (unfold list_less_def) simp
 | 
| 15737 |     46 | 
 | 
| 17200 |     47 | lemma le_Nil[simp]: "(x <= []) = (x = [])"
 | 
|  |     48 |   by (unfold list_ord_defs, cases x) auto
 | 
| 15737 |     49 | 
 | 
| 17200 |     50 | lemma Nil_le_Cons [simp]: "([] <= x)"
 | 
|  |     51 |   by (unfold list_ord_defs, cases x) auto
 | 
| 15737 |     52 | 
 | 
| 17200 |     53 | lemma Cons_le_Cons[simp]: "(a # x <= b # y) = (a < b | a = b & x <= y)"
 | 
|  |     54 |   by (unfold list_ord_defs) auto
 | 
| 15737 |     55 | 
 | 
| 17200 |     56 | end
 |