nipkow@15737: (* Title: HOL/Library/List_lexord.thy nipkow@15737: Author: Norbert Voelker nipkow@15737: *) nipkow@15737: wenzelm@17200: header {* Lexicographic order on lists *} nipkow@15737: nipkow@15737: theory List_lexord haftmann@30663: imports List Main nipkow@15737: begin nipkow@15737: haftmann@25764: instantiation list :: (ord) ord haftmann@25764: begin haftmann@25764: haftmann@25764: definition haftmann@37474: list_less_def: "xs < ys \ (xs, ys) \ lexord {(u, v). u < v}" haftmann@25764: haftmann@25764: definition haftmann@37474: list_le_def: "(xs :: _ list) \ ys \ xs < ys \ xs = ys" haftmann@25764: haftmann@25764: instance .. haftmann@25764: haftmann@25764: end nipkow@15737: wenzelm@17200: instance list :: (order) order haftmann@27682: proof haftmann@27682: fix xs :: "'a list" haftmann@27682: show "xs \ xs" by (simp add: list_le_def) haftmann@27682: next haftmann@27682: fix xs ys zs :: "'a list" haftmann@27682: assume "xs \ ys" and "ys \ zs" haftmann@27682: then show "xs \ zs" by (auto simp add: list_le_def list_less_def) haftmann@27682: (rule lexord_trans, auto intro: transI) haftmann@27682: next haftmann@27682: fix xs ys :: "'a list" haftmann@27682: assume "xs \ ys" and "ys \ xs" haftmann@27682: then show "xs = ys" apply (auto simp add: list_le_def list_less_def) haftmann@27682: apply (rule lexord_irreflexive [THEN notE]) haftmann@27682: defer haftmann@27682: apply (rule lexord_trans) apply (auto intro: transI) done haftmann@27682: next haftmann@27682: fix xs ys :: "'a list" haftmann@27682: show "xs < ys \ xs \ ys \ \ ys \ xs" haftmann@27682: apply (auto simp add: list_less_def list_le_def) haftmann@27682: defer haftmann@27682: apply (rule lexord_irreflexive [THEN notE]) haftmann@27682: apply auto haftmann@27682: apply (rule lexord_irreflexive [THEN notE]) haftmann@27682: defer haftmann@27682: apply (rule lexord_trans) apply (auto intro: transI) done haftmann@27682: qed nipkow@15737: haftmann@21458: instance list :: (linorder) linorder haftmann@27682: proof haftmann@27682: fix xs ys :: "'a list" haftmann@27682: have "(xs, ys) \ lexord {(u, v). u < v} \ xs = ys \ (ys, xs) \ lexord {(u, v). u < v}" haftmann@27682: by (rule lexord_linear) auto haftmann@27682: then show "xs \ ys \ ys \ xs" haftmann@27682: by (auto simp add: list_le_def list_less_def) haftmann@27682: qed nipkow@15737: haftmann@25764: instantiation list :: (linorder) distrib_lattice haftmann@25764: begin haftmann@25764: haftmann@25764: definition haftmann@37765: "(inf \ 'a list \ _) = min" haftmann@25764: haftmann@25764: definition haftmann@37765: "(sup \ 'a list \ _) = max" haftmann@25764: haftmann@25764: instance haftmann@22483: by intro_classes haftmann@22483: (auto simp add: inf_list_def sup_list_def min_max.sup_inf_distrib1) haftmann@22483: haftmann@25764: end haftmann@25764: haftmann@22177: lemma not_less_Nil [simp]: "\ (x < [])" wenzelm@17200: by (unfold list_less_def) simp nipkow@15737: haftmann@22177: lemma Nil_less_Cons [simp]: "[] < a # x" wenzelm@17200: by (unfold list_less_def) simp nipkow@15737: haftmann@22177: lemma Cons_less_Cons [simp]: "a # x < b # y \ a < b \ a = b \ x < y" wenzelm@17200: by (unfold list_less_def) simp nipkow@15737: haftmann@22177: lemma le_Nil [simp]: "x \ [] \ x = []" haftmann@25502: by (unfold list_le_def, cases x) auto haftmann@22177: haftmann@22177: lemma Nil_le_Cons [simp]: "[] \ x" haftmann@25502: by (unfold list_le_def, cases x) auto nipkow@15737: haftmann@22177: lemma Cons_le_Cons [simp]: "a # x \ b # y \ a < b \ a = b \ x \ y" haftmann@25502: by (unfold list_le_def) auto nipkow@15737: haftmann@37474: instantiation list :: (order) bot haftmann@37474: begin haftmann@37474: haftmann@37474: definition haftmann@37474: "bot = []" haftmann@37474: haftmann@37474: instance proof haftmann@37474: qed (simp add: bot_list_def) haftmann@37474: haftmann@37474: end haftmann@37474: haftmann@37474: lemma less_list_code [code]: haftmann@38857: "xs < ([]\'a\{equal, order} list) \ False" haftmann@38857: "[] < (x\'a\{equal, order}) # xs \ True" haftmann@38857: "(x\'a\{equal, order}) # xs < y # ys \ x < y \ x = y \ xs < ys" haftmann@22177: by simp_all haftmann@22177: haftmann@37474: lemma less_eq_list_code [code]: haftmann@38857: "x # xs \ ([]\'a\{equal, order} list) \ False" haftmann@38857: "[] \ (xs\'a\{equal, order} list) \ True" haftmann@38857: "(x\'a\{equal, order}) # xs \ y # ys \ x < y \ x = y \ xs \ ys" haftmann@22177: by simp_all haftmann@22177: wenzelm@17200: end