# HG changeset patch # User paulson # Date 1587160766 -3600 # Node ID ed065b4a85f2da41245516a822438dbcc06a1a30 # Parent c4e0d29cf1ca4f65181235163b67d0956c7a2a9d# Parent 1249b998e3779a4b9294f9c8ae1593184a166907 merged diff -r c4e0d29cf1ca -r ed065b4a85f2 src/HOL/Library/List_Lenlexorder.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/List_Lenlexorder.thy Fri Apr 17 22:59:26 2020 +0100 @@ -0,0 +1,95 @@ +(* Title: HOL/Library/List_Lenlexorder.thy +*) + +section \Lexicographic order on lists\ + +theory List_Lenlexorder +imports Main +begin + + +instantiation list :: (ord) ord +begin + +definition + list_less_def: "xs < ys \ (xs, ys) \ lenlex {(u, v). u < v}" + +definition + list_le_def: "(xs :: _ list) \ ys \ xs < ys \ xs = ys" + +instance .. + +end + +instance list :: (order) order +proof + have tr: "trans {(u, v::'a). u < v}" + using trans_def by fastforce + have \
: False + if "(xs,ys) \ lenlex {(u, v). u < v}" "(ys,xs) \ lenlex {(u, v). u < v}" for xs ys :: "'a list" + proof - + have "(xs,xs) \ lenlex {(u, v). u < v}" + using that transD [OF lenlex_transI [OF tr]] by blast + then show False + by (meson case_prodD lenlex_irreflexive less_irrefl mem_Collect_eq) + qed + show "xs \ xs" for xs :: "'a list" by (simp add: list_le_def) + show "xs \ zs" if "xs \ ys" and "ys \ zs" for xs ys zs :: "'a list" + using that transD [OF lenlex_transI [OF tr]] by (auto simp add: list_le_def list_less_def) + show "xs = ys" if "xs \ ys" "ys \ xs" for xs ys :: "'a list" + using \
that list_le_def list_less_def by blast + show "xs < ys \ xs \ ys \ \ ys \ xs" for xs ys :: "'a list" + by (auto simp add: list_less_def list_le_def dest: \
) +qed + +instance list :: (linorder) linorder +proof + fix xs ys :: "'a list" + have "total (lenlex {(u, v::'a). u < v})" + by (rule total_lenlex) (auto simp: total_on_def) + then show "xs \ ys \ ys \ xs" + by (auto simp add: total_on_def 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: "a # x < b # y \ length x < length y \ length x = length y \ (a < b \ a = b \ x < y)" + using lenlex_length + by (fastforce simp: list_less_def Cons_lenlex_iff) + +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: "a # x \ b # y \ length x < length y \ length x = length y \ (a < b \ a = b \ x \ y)" + by (auto simp: list_le_def Cons_less_Cons) + +instantiation list :: (order) order_bot +begin + +definition "bot = []" + +instance + by standard (simp add: bot_list_def) + +end + +end diff -r c4e0d29cf1ca -r ed065b4a85f2 src/HOL/Library/List_Lexorder.thy --- a/src/HOL/Library/List_Lexorder.thy Fri Apr 17 22:13:50 2020 +0200 +++ b/src/HOL/Library/List_Lexorder.thy Fri Apr 17 22:59:26 2020 +0100 @@ -23,47 +23,33 @@ 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 + let ?r = "{(u, v::'a). u < v}" + have tr: "trans ?r" + using trans_def by fastforce + have \
: False + if "(xs,ys) \ lexord ?r" "(ys,xs) \ lexord ?r" for xs ys :: "'a list" + proof - + have "(xs,xs) \ lexord ?r" + using that transD [OF lexord_transI [OF tr]] by blast + then show False + by (meson case_prodD lexord_irreflexive less_irrefl mem_Collect_eq) + qed + show "xs \ xs" for xs :: "'a list" by (simp add: list_le_def) + show "xs \ zs" if "xs \ ys" and "ys \ zs" for xs ys zs :: "'a list" + using that transD [OF lexord_transI [OF tr]] by (auto simp add: list_le_def list_less_def) + show "xs = ys" if "xs \ ys" "ys \ xs" for xs ys :: "'a list" + using \
that list_le_def list_less_def by blast + show "xs < ys \ xs \ ys \ \ ys \ xs" for xs ys :: "'a list" + by (auto simp add: list_less_def list_le_def dest: \
) 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 + have "total (lexord {(u, v::'a). u < v})" + by (rule total_lexord) (auto simp: total_on_def) then show "xs \ ys \ ys \ xs" - by (auto simp add: list_le_def list_less_def) + by (auto simp add: total_on_def list_le_def list_less_def) qed instantiation list :: (linorder) distrib_lattice diff -r c4e0d29cf1ca -r ed065b4a85f2 src/HOL/List.thy --- a/src/HOL/List.thy Fri Apr 17 22:13:50 2020 +0200 +++ b/src/HOL/List.thy Fri Apr 17 22:59:26 2020 +0100 @@ -6178,6 +6178,13 @@ \ (m = n \ (ms,ns) \ lenlex r)" by (auto simp: lenlex_def) +lemma lenlex_irreflexive: "(\x. (x,x) \ r) \ (xs,xs) \ lenlex r" + by (induction xs) (auto simp add: Cons_lenlex_iff) + +lemma lenlex_trans: + "\(x,y) \ lenlex r; (y,z) \ lenlex r; trans r\ \ (x,z) \ lenlex r" + by (meson lenlex_transI transD) + lemma lenlex_length: "(ms, ns) \ lenlex r \ length ms \ length ns" by (auto simp: lenlex_def) @@ -6316,15 +6323,27 @@ lemma lexord_transI: "trans r \ trans (lexord r)" by (rule transI, drule lexord_trans, blast) -lemma lexord_linear: "(\a b. (a,b) \ r \ a = b \ (b,a) \ r) \ (x,y) \ lexord r \ x = y \ (y,x) \ lexord r" -proof (induction x arbitrary: y) - case Nil - then show ?case - by (metis lexord_Nil_left list.exhaust) -next - case (Cons a x y) then show ?case - by (cases y) (force+) -qed +lemma total_lexord: "total r \ total (lexord r)" + unfolding total_on_def +proof clarsimp + fix x y + assume "\x y. x \ y \ (x, y) \ r \ (y, x) \ r" + and "(x::'a list) \ y" + and "(y, x) \ lexord r" + then + show "(x, y) \ lexord r" + proof (induction x arbitrary: y) + case Nil + then show ?case + by (metis lexord_Nil_left list.exhaust) + next + case (Cons a x y) then show ?case + by (cases y) (force+) + qed +qed + +corollary lexord_linear: "(\a b. (a,b) \ r \ a = b \ (b,a) \ r) \ (x,y) \ lexord r \ x = y \ (y,x) \ lexord r" + using total_lexord by (metis UNIV_I total_on_def) lemma lexord_irrefl: "irrefl R \ irrefl (lexord R)" diff -r c4e0d29cf1ca -r ed065b4a85f2 src/HOL/ROOT --- a/src/HOL/ROOT Fri Apr 17 22:13:50 2020 +0200 +++ b/src/HOL/ROOT Fri Apr 17 22:59:26 2020 +0100 @@ -32,6 +32,7 @@ (*conflicting type class instantiations and dependent applications*) Finite_Lattice List_Lexorder + List_Lenlexorder Prefix_Order Product_Lexorder Product_Order diff -r c4e0d29cf1ca -r ed065b4a85f2 src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Fri Apr 17 22:13:50 2020 +0200 +++ b/src/HOL/Wellfounded.thy Fri Apr 17 22:59:26 2020 +0100 @@ -580,8 +580,8 @@ lemma less_than_iff [iff]: "((x,y) \ less_than) = (x