| author | wenzelm | 
| Sat, 13 Aug 2022 14:45:36 +0200 | |
| changeset 75841 | 7c00d5266bf8 | 
| parent 75716 | f6695e7aff32 | 
| permissions | -rw-r--r-- | 
| 71766 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | (* Title: HOL/Library/List_Lenlexorder.thy | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | *) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | section \<open>Lexicographic order on lists\<close> | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | |
| 75716 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 6 | text \<open>This version prioritises length and can yield wellorderings\<close> | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 7 | |
| 71766 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | theory List_Lenlexorder | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | imports Main | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | begin | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | instantiation list :: (ord) ord | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | begin | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | definition | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 |   list_less_def: "xs < ys \<longleftrightarrow> (xs, ys) \<in> lenlex {(u, v). u < v}"
 | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | definition | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | list_le_def: "(xs :: _ list) \<le> ys \<longleftrightarrow> xs < ys \<or> xs = ys" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | instance .. | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | end | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | instance list :: (order) order | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | proof | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 |   have tr: "trans {(u, v::'a). u < v}"
 | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | using trans_def by fastforce | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | have \<section>: False | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 |     if "(xs,ys) \<in> lenlex {(u, v). u < v}" "(ys,xs) \<in> lenlex {(u, v). u < v}" for xs ys :: "'a list"
 | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | proof - | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 |     have "(xs,xs) \<in> lenlex {(u, v). u < v}"
 | 
| 72184 | 34 | using that transD [OF lenlex_transI [OF tr]] by blast | 
| 71766 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | then show False | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | by (meson case_prodD lenlex_irreflexive less_irrefl mem_Collect_eq) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | qed | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | show "xs \<le> xs" for xs :: "'a list" by (simp add: list_le_def) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | show "xs \<le> zs" if "xs \<le> ys" and "ys \<le> zs" for xs ys zs :: "'a list" | 
| 72184 | 40 | using that transD [OF lenlex_transI [OF tr]] by (auto simp add: list_le_def list_less_def) | 
| 71766 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | show "xs = ys" if "xs \<le> ys" "ys \<le> xs" for xs ys :: "'a list" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | using \<section> that list_le_def list_less_def by blast | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | show "xs < ys \<longleftrightarrow> xs \<le> ys \<and> \<not> ys \<le> xs" for xs ys :: "'a list" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | by (auto simp add: list_less_def list_le_def dest: \<section>) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | qed | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | instance list :: (linorder) linorder | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | proof | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | fix xs ys :: "'a list" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 |   have "total (lenlex {(u, v::'a). u < v})"
 | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | by (rule total_lenlex) (auto simp: total_on_def) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | then show "xs \<le> ys \<or> ys \<le> xs" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | by (auto simp add: total_on_def list_le_def list_less_def) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | qed | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | |
| 75716 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 56 | instance list :: (wellorder) wellorder | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 57 | proof | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 58 | fix P :: "'a list \<Rightarrow> bool" and a | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 59 | assume "\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x" | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 60 | then show "P a" | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 61 | unfolding list_less_def by (metis wf_lenlex wf_induct wf_lenlex wf) | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 62 | qed | 
| 
f6695e7aff32
The wellordering instantiation for length-ordered lists
 paulson <lp15@cam.ac.uk> parents: 
72184diff
changeset | 63 | |
| 71766 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | instantiation list :: (linorder) distrib_lattice | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | begin | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | definition "(inf :: 'a list \<Rightarrow> _) = min" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | definition "(sup :: 'a list \<Rightarrow> _) = max" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | instance | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | by standard (auto simp add: inf_list_def sup_list_def max_min_distrib2) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | end | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | lemma not_less_Nil [simp]: "\<not> x < []" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | by (simp add: list_less_def) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | lemma Nil_less_Cons [simp]: "[] < a # x" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | by (simp add: list_less_def) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | |
| 72184 | 82 | lemma Cons_less_Cons: "a # x < b # y \<longleftrightarrow> length x < length y \<or> length x = length y \<and> (a < b \<or> a = b \<and> x < y)" | 
| 83 | using lenlex_length | |
| 84 | by (fastforce simp: list_less_def Cons_lenlex_iff) | |
| 85 | ||
| 71766 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | lemma le_Nil [simp]: "x \<le> [] \<longleftrightarrow> x = []" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | unfolding list_le_def by (cases x) auto | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | lemma Nil_le_Cons [simp]: "[] \<le> x" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | unfolding list_le_def by (cases x) auto | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | |
| 72184 | 92 | lemma Cons_le_Cons: "a # x \<le> b # y \<longleftrightarrow> length x < length y \<or> length x = length y \<and> (a < b \<or> a = b \<and> x \<le> y)" | 
| 93 | by (auto simp: list_le_def Cons_less_Cons) | |
| 94 | ||
| 71766 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | instantiation list :: (order) order_bot | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | begin | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | definition "bot = []" | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | instance | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | by standard (simp add: bot_list_def) | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | end | 
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | |
| 
1249b998e377
New theory Library/List_Lenlexorder.thy, a type class instantiation for well-ordering lists
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | end |