# HG changeset patch # User nipkow # Date 1527595559 -7200 # Node ID e9b5f25f6712aaa9016cc85e43d2fbbbd6e98932 # Parent c551d8acaa4206bc3d3ddee9935d53cca285bd03 canonical names diff -r c551d8acaa42 -r e9b5f25f6712 src/HOL/Library/Fun_Lexorder.thy --- a/src/HOL/Library/Fun_Lexorder.thy Mon May 28 23:15:30 2018 +0100 +++ b/src/HOL/Library/Fun_Lexorder.thy Tue May 29 14:05:59 2018 +0200 @@ -1,6 +1,6 @@ (* Author: Florian Haftmann, TU Muenchen *) -section \Lexical order on functions\ +section \Lexicographic order on functions\ theory Fun_Lexorder imports Main diff -r c551d8acaa42 -r e9b5f25f6712 src/HOL/Library/List_Lexorder.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/List_Lexorder.thy Tue May 29 14:05:59 2018 +0200 @@ -0,0 +1,121 @@ +(* Title: HOL/Library/List_Lexorder.thy + Author: Norbert Voelker +*) + +section \Lexicographic order on lists\ + +theory List_Lexorder +imports Main +begin + +instantiation list :: (ord) ord +begin + +definition + list_less_def: "xs < ys \ (xs, ys) \ lexord {(u, v). u < v}" + +definition + list_le_def: "(xs :: _ list) \ ys \ xs < ys \ xs = ys" + +instance .. + +end + +instance list :: (order) order +proof + fix xs :: "'a list" + show "xs \ xs" by (simp add: list_le_def) +next + fix xs ys zs :: "'a list" + assume "xs \ ys" and "ys \ zs" + 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 +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 +qed + +instance list :: (linorder) linorder +proof + 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" + by (auto simp add: list_le_def list_less_def) +qed + +instantiation list :: (linorder) distrib_lattice +begin + +definition "(inf :: 'a list \ _) = min" + +definition "(sup :: 'a list \ _) = max" + +instance + by standard (auto simp add: inf_list_def sup_list_def max_min_distrib2) + +end + +lemma not_less_Nil [simp]: "\ x < []" + by (simp add: list_less_def) + +lemma Nil_less_Cons [simp]: "[] < a # x" + by (simp add: list_less_def) + +lemma Cons_less_Cons [simp]: "a # x < b # y \ a < b \ a = b \ x < y" + by (simp add: list_less_def) + +lemma le_Nil [simp]: "x \ [] \ x = []" + unfolding list_le_def by (cases x) auto + +lemma Nil_le_Cons [simp]: "[] \ x" + unfolding list_le_def by (cases x) auto + +lemma Cons_le_Cons [simp]: "a # x \ b # y \ a < b \ a = b \ x \ y" + unfolding list_le_def by auto + +instantiation list :: (order) order_bot +begin + +definition "bot = []" + +instance + by standard (simp add: bot_list_def) + +end + +lemma less_list_code [code]: + "xs < ([]::'a::{equal, order} list) \ False" + "[] < (x::'a::{equal, order}) # xs \ True" + "(x::'a::{equal, order}) # xs < y # ys \ x < y \ x = y \ xs < ys" + by simp_all + +lemma less_eq_list_code [code]: + "x # xs \ ([]::'a::{equal, order} list) \ False" + "[] \ (xs::'a::{equal, order} list) \ True" + "(x::'a::{equal, order}) # xs \ y # ys \ x < y \ x = y \ xs \ ys" + by simp_all + +end diff -r c551d8acaa42 -r e9b5f25f6712 src/HOL/Library/List_lexord.thy --- a/src/HOL/Library/List_lexord.thy Mon May 28 23:15:30 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,121 +0,0 @@ -(* Title: HOL/Library/List_lexord.thy - Author: Norbert Voelker -*) - -section \Lexicographic order on lists\ - -theory List_lexord -imports Main -begin - -instantiation list :: (ord) ord -begin - -definition - list_less_def: "xs < ys \ (xs, ys) \ lexord {(u, v). u < v}" - -definition - list_le_def: "(xs :: _ list) \ ys \ xs < ys \ xs = ys" - -instance .. - -end - -instance list :: (order) order -proof - fix xs :: "'a list" - show "xs \ xs" by (simp add: list_le_def) -next - fix xs ys zs :: "'a list" - assume "xs \ ys" and "ys \ zs" - 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 -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 -qed - -instance list :: (linorder) linorder -proof - 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" - by (auto simp add: list_le_def list_less_def) -qed - -instantiation list :: (linorder) distrib_lattice -begin - -definition "(inf :: 'a list \ _) = min" - -definition "(sup :: 'a list \ _) = max" - -instance - by standard (auto simp add: inf_list_def sup_list_def max_min_distrib2) - -end - -lemma not_less_Nil [simp]: "\ x < []" - by (simp add: list_less_def) - -lemma Nil_less_Cons [simp]: "[] < a # x" - by (simp add: list_less_def) - -lemma Cons_less_Cons [simp]: "a # x < b # y \ a < b \ a = b \ x < y" - by (simp add: list_less_def) - -lemma le_Nil [simp]: "x \ [] \ x = []" - unfolding list_le_def by (cases x) auto - -lemma Nil_le_Cons [simp]: "[] \ x" - unfolding list_le_def by (cases x) auto - -lemma Cons_le_Cons [simp]: "a # x \ b # y \ a < b \ a = b \ x \ y" - unfolding list_le_def by auto - -instantiation list :: (order) order_bot -begin - -definition "bot = []" - -instance - by standard (simp add: bot_list_def) - -end - -lemma less_list_code [code]: - "xs < ([]::'a::{equal, order} list) \ False" - "[] < (x::'a::{equal, order}) # xs \ True" - "(x::'a::{equal, order}) # xs < y # ys \ x < y \ x = y \ xs < ys" - by simp_all - -lemma less_eq_list_code [code]: - "x # xs \ ([]::'a::{equal, order} list) \ False" - "[] \ (xs::'a::{equal, order} list) \ True" - "(x::'a::{equal, order}) # xs \ y # ys \ x < y \ x = y \ xs \ ys" - by simp_all - -end diff -r c551d8acaa42 -r e9b5f25f6712 src/HOL/ROOT --- a/src/HOL/ROOT Mon May 28 23:15:30 2018 +0100 +++ b/src/HOL/ROOT Tue May 29 14:05:59 2018 +0200 @@ -29,7 +29,7 @@ Library (*conflicting type class instantiations and dependent applications*) Finite_Lattice - List_lexord + List_Lexorder Prefix_Order Product_Lexorder Product_Order diff -r c551d8acaa42 -r e9b5f25f6712 src/HOL/ex/Radix_Sort.thy --- a/src/HOL/ex/Radix_Sort.thy Mon May 28 23:15:30 2018 +0100 +++ b/src/HOL/ex/Radix_Sort.thy Tue May 29 14:05:59 2018 +0200 @@ -2,7 +2,7 @@ theory Radix_Sort imports - "HOL-Library.List_lexord" + "HOL-Library.List_Lexorder" "HOL-Library.Sublist" "HOL-Library.Multiset" begin