# HG changeset patch # User wenzelm # Date 1377553193 -7200 # Node ID bae01293f4dd2bcb0d04a1399e14eee08c211388 # Parent a11e55f667dbcb00b29a49c47a5143baea3a2ef3 tuned imports; tuned proofs; diff -r a11e55f667db -r bae01293f4dd src/HOL/Library/List_lexord.thy --- a/src/HOL/Library/List_lexord.thy Mon Aug 26 22:01:39 2013 +0200 +++ b/src/HOL/Library/List_lexord.thy Mon Aug 26 23:39:53 2013 +0200 @@ -5,7 +5,7 @@ header {* Lexicographic order on lists *} theory List_lexord -imports List Main +imports Main begin instantiation list :: (ord) ord @@ -28,25 +28,33 @@ next fix xs ys zs :: "'a list" assume "xs \ ys" and "ys \ zs" - then show "xs \ zs" by (auto simp add: list_le_def list_less_def) - (rule lexord_trans, auto intro: transI) + then show "xs \ zs" + apply (auto simp add: list_le_def list_less_def) + apply (rule lexord_trans) + apply (auto intro: transI) + done next fix xs ys :: "'a list" assume "xs \ ys" and "ys \ xs" - then show "xs = ys" apply (auto simp add: list_le_def list_less_def) - apply (rule lexord_irreflexive [THEN notE]) - defer - apply (rule lexord_trans) apply (auto intro: transI) done + then show "xs = ys" + apply (auto simp add: list_le_def list_less_def) + apply (rule lexord_irreflexive [THEN notE]) + defer + apply (rule lexord_trans) + apply (auto intro: transI) + done next fix xs ys :: "'a list" - show "xs < ys \ xs \ ys \ \ ys \ xs" - apply (auto simp add: list_less_def list_le_def) - defer - apply (rule lexord_irreflexive [THEN notE]) - apply auto - apply (rule lexord_irreflexive [THEN notE]) - defer - apply (rule lexord_trans) apply (auto intro: transI) done + show "xs < ys \ xs \ ys \ \ ys \ xs" + apply (auto simp add: list_less_def list_le_def) + defer + apply (rule lexord_irreflexive [THEN notE]) + apply auto + apply (rule lexord_irreflexive [THEN notE]) + defer + apply (rule lexord_trans) + apply (auto intro: transI) + done qed instance list :: (linorder) linorder @@ -54,51 +62,47 @@ fix xs ys :: "'a list" have "(xs, ys) \ lexord {(u, v). u < v} \ xs = ys \ (ys, xs) \ lexord {(u, v). u < v}" by (rule lexord_linear) auto - then show "xs \ ys \ ys \ xs" + then show "xs \ ys \ ys \ xs" by (auto simp add: list_le_def list_less_def) qed instantiation list :: (linorder) distrib_lattice begin -definition - "(inf \ 'a list \ _) = min" +definition "(inf \ 'a list \ _) = min" -definition - "(sup \ 'a list \ _) = max" +definition "(sup \ 'a list \ _) = max" instance - by intro_classes - (auto simp add: inf_list_def sup_list_def min_max.sup_inf_distrib1) + by default (auto simp add: inf_list_def sup_list_def min_max.sup_inf_distrib1) end -lemma not_less_Nil [simp]: "\ (x < [])" - by (unfold list_less_def) simp +lemma not_less_Nil [simp]: "\ x < []" + by (simp add: list_less_def) lemma Nil_less_Cons [simp]: "[] < a # x" - by (unfold list_less_def) simp + by (simp add: list_less_def) lemma Cons_less_Cons [simp]: "a # x < b # y \ a < b \ a = b \ x < y" - by (unfold list_less_def) simp + by (simp add: list_less_def) lemma le_Nil [simp]: "x \ [] \ x = []" - by (unfold list_le_def, cases x) auto + unfolding list_le_def by (cases x) auto lemma Nil_le_Cons [simp]: "[] \ x" - by (unfold list_le_def, cases x) auto + unfolding list_le_def by (cases x) auto lemma Cons_le_Cons [simp]: "a # x \ b # y \ a < b \ a = b \ x \ y" - by (unfold list_le_def) auto + unfolding list_le_def by auto instantiation list :: (order) order_bot begin -definition - "bot = []" +definition "bot = []" -instance proof -qed (simp add: bot_list_def) +instance + by default (simp add: bot_list_def) end