| author | wenzelm | 
| Fri, 27 Sep 2024 22:44:30 +0200 | |
| changeset 80982 | 7638776e0977 | 
| parent 77955 | c4677a6aae2c | 
| child 82774 | 2865a6618cba | 
| permissions | -rw-r--r-- | 
| 73794 | 1 | section \<open>Lexicographic orderings\<close> | 
| 2 | ||
| 3 | theory Lexord | |
| 4 | imports Main | |
| 5 | begin | |
| 6 | ||
| 7 | subsection \<open>The preorder case\<close> | |
| 8 | ||
| 9 | locale lex_preordering = preordering | |
| 10 | begin | |
| 11 | ||
| 12 | inductive lex_less :: \<open>'a list \<Rightarrow> 'a list \<Rightarrow> bool\<close> (infix \<open>[\<^bold><]\<close> 50) | |
| 13 | where | |
| 14 | Nil: \<open>[] [\<^bold><] y # ys\<close> | |
| 15 | | Cons: \<open>x \<^bold>< y \<Longrightarrow> x # xs [\<^bold><] y # ys\<close> | |
| 16 | | Cons_eq: \<open>x \<^bold>\<le> y \<Longrightarrow> y \<^bold>\<le> x \<Longrightarrow> xs [\<^bold><] ys \<Longrightarrow> x # xs [\<^bold><] y # ys\<close> | |
| 17 | ||
| 18 | inductive lex_less_eq :: \<open>'a list \<Rightarrow> 'a list \<Rightarrow> bool\<close> (infix \<open>[\<^bold>\<le>]\<close> 50) | |
| 19 | where | |
| 20 | Nil: \<open>[] [\<^bold>\<le>] ys\<close> | |
| 21 | | Cons: \<open>x \<^bold>< y \<Longrightarrow> x # xs [\<^bold>\<le>] y # ys\<close> | |
| 22 | | Cons_eq: \<open>x \<^bold>\<le> y \<Longrightarrow> y \<^bold>\<le> x \<Longrightarrow> xs [\<^bold>\<le>] ys \<Longrightarrow> x # xs [\<^bold>\<le>] y # ys\<close> | |
| 23 | ||
| 24 | lemma lex_less_simps [simp]: | |
| 25 | \<open>[] [\<^bold><] y # ys\<close> | |
| 26 | \<open>\<not> xs [\<^bold><] []\<close> | |
| 27 | \<open>x # xs [\<^bold><] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold><] ys\<close> | |
| 28 | by (auto intro: lex_less.intros elim: lex_less.cases) | |
| 29 | ||
| 30 | lemma lex_less_eq_simps [simp]: | |
| 31 | \<open>[] [\<^bold>\<le>] ys\<close> | |
| 32 | \<open>\<not> x # xs [\<^bold>\<le>] []\<close> | |
| 33 | \<open>x # xs [\<^bold>\<le>] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold>\<le>] ys\<close> | |
| 34 | by (auto intro: lex_less_eq.intros elim: lex_less_eq.cases) | |
| 35 | ||
| 36 | lemma lex_less_code [code]: | |
| 37 | \<open>[] [\<^bold><] y # ys \<longleftrightarrow> True\<close> | |
| 38 | \<open>xs [\<^bold><] [] \<longleftrightarrow> False\<close> | |
| 39 | \<open>x # xs [\<^bold><] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold><] ys\<close> | |
| 40 | by simp_all | |
| 41 | ||
| 42 | lemma lex_less_eq_code [code]: | |
| 43 | \<open>[] [\<^bold>\<le>] ys \<longleftrightarrow> True\<close> | |
| 44 | \<open>x # xs [\<^bold>\<le>] [] \<longleftrightarrow> False\<close> | |
| 45 | \<open>x # xs [\<^bold>\<le>] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold>\<le>] ys\<close> | |
| 46 | by simp_all | |
| 47 | ||
| 48 | lemma preordering: | |
| 49 | \<open>preordering ([\<^bold>\<le>]) ([\<^bold><])\<close> | |
| 50 | proof | |
| 51 | fix xs ys zs | |
| 52 | show \<open>xs [\<^bold>\<le>] xs\<close> | |
| 53 | by (induction xs) (simp_all add: refl) | |
| 54 | show \<open>xs [\<^bold>\<le>] zs\<close> if \<open>xs [\<^bold>\<le>] ys\<close> \<open>ys [\<^bold>\<le>] zs\<close> | |
| 55 | using that proof (induction arbitrary: zs) | |
| 56 | case (Nil ys) | |
| 57 | then show ?case by simp | |
| 58 | next | |
| 59 | case (Cons x y xs ys) | |
| 60 | then show ?case | |
| 61 | by (cases zs) (auto dest: strict_trans strict_trans2) | |
| 62 | next | |
| 63 | case (Cons_eq x y xs ys) | |
| 64 | then show ?case | |
| 65 | by (cases zs) (auto dest: strict_trans1 intro: trans) | |
| 66 | qed | |
| 67 | show \<open>xs [\<^bold><] ys \<longleftrightarrow> xs [\<^bold>\<le>] ys \<and> \<not> ys [\<^bold>\<le>] xs\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 68 | proof | |
| 69 | assume ?P | |
| 70 | then have \<open>xs [\<^bold>\<le>] ys\<close> | |
| 71 | by induction simp_all | |
| 72 | moreover have \<open>\<not> ys [\<^bold>\<le>] xs\<close> | |
| 73 | using \<open>?P\<close> | |
| 74 | by induction (simp_all, simp_all add: strict_iff_not asym) | |
| 75 | ultimately show ?Q .. | |
| 76 | next | |
| 77 | assume ?Q | |
| 78 | then have \<open>xs [\<^bold>\<le>] ys\<close> \<open>\<not> ys [\<^bold>\<le>] xs\<close> | |
| 79 | by auto | |
| 80 | then show ?P | |
| 81 | proof induction | |
| 82 | case (Nil ys) | |
| 83 | then show ?case | |
| 84 | by (cases ys) simp_all | |
| 85 | next | |
| 86 | case (Cons x y xs ys) | |
| 87 | then show ?case | |
| 88 | by simp | |
| 89 | next | |
| 90 | case (Cons_eq x y xs ys) | |
| 91 | then show ?case | |
| 92 | by simp | |
| 93 | qed | |
| 94 | qed | |
| 95 | qed | |
| 96 | ||
| 97 | interpretation lex: preordering \<open>([\<^bold>\<le>])\<close> \<open>([\<^bold><])\<close> | |
| 98 | by (fact preordering) | |
| 99 | ||
| 100 | end | |
| 101 | ||
| 102 | ||
| 103 | subsection \<open>The order case\<close> | |
| 104 | ||
| 105 | locale lex_ordering = lex_preordering + ordering | |
| 106 | begin | |
| 107 | ||
| 108 | interpretation lex: preordering \<open>([\<^bold>\<le>])\<close> \<open>([\<^bold><])\<close> | |
| 109 | by (fact preordering) | |
| 110 | ||
| 111 | lemma less_lex_Cons_iff [simp]: | |
| 112 | \<open>x # xs [\<^bold><] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x = y \<and> xs [\<^bold><] ys\<close> | |
| 113 | by (auto intro: refl antisym) | |
| 114 | ||
| 115 | lemma less_eq_lex_Cons_iff [simp]: | |
| 116 | \<open>x # xs [\<^bold>\<le>] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x = y \<and> xs [\<^bold>\<le>] ys\<close> | |
| 117 | by (auto intro: refl antisym) | |
| 118 | ||
| 119 | lemma ordering: | |
| 120 | \<open>ordering ([\<^bold>\<le>]) ([\<^bold><])\<close> | |
| 121 | proof | |
| 122 | fix xs ys | |
| 123 | show *: \<open>xs = ys\<close> if \<open>xs [\<^bold>\<le>] ys\<close> \<open>ys [\<^bold>\<le>] xs\<close> | |
| 124 | using that proof induction | |
| 125 | case (Nil ys) | |
| 126 | then show ?case by (cases ys) simp | |
| 127 | next | |
| 128 | case (Cons x y xs ys) | |
| 129 | then show ?case by (auto dest: asym intro: antisym) | |
| 130 | (simp add: strict_iff_not) | |
| 131 | next | |
| 132 | case (Cons_eq x y xs ys) | |
| 133 | then show ?case by (auto intro: antisym) | |
| 134 | (simp add: strict_iff_not) | |
| 135 | qed | |
| 136 | show \<open>xs [\<^bold><] ys \<longleftrightarrow> xs [\<^bold>\<le>] ys \<and> xs \<noteq> ys\<close> | |
| 137 | by (auto simp add: lex.strict_iff_not dest: *) | |
| 138 | qed | |
| 139 | ||
| 140 | interpretation lex: ordering \<open>([\<^bold>\<le>])\<close> \<open>([\<^bold><])\<close> | |
| 141 | by (fact ordering) | |
| 142 | ||
| 143 | end | |
| 144 | ||
| 145 | ||
| 146 | subsection \<open>Canonical instance\<close> | |
| 147 | ||
| 148 | instantiation list :: (preorder) preorder | |
| 149 | begin | |
| 150 | ||
| 151 | global_interpretation lex: lex_preordering \<open>(\<le>) :: 'a::preorder \<Rightarrow> 'a \<Rightarrow> bool\<close> \<open>(<) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close> | |
| 152 | defines less_eq_list = lex.lex_less_eq | |
| 153 | and less_list = lex.lex_less .. | |
| 154 | ||
| 155 | instance | |
| 77955 
c4677a6aae2c
more standard name bindings (amending 5bf71b4da706): avoid odd full_name like "Orderings.class.Orderings.preorder.of_class.intro" with many redundant name space accesses;
 wenzelm parents: 
75244diff
changeset | 156 | by (rule preorder.intro_of_class, rule preordering_preorderI, fact lex.preordering) | 
| 73794 | 157 | |
| 158 | end | |
| 159 | ||
| 160 | global_interpretation lex: lex_ordering \<open>(\<le>) :: 'a::order \<Rightarrow> 'a \<Rightarrow> bool\<close> \<open>(<) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close> | |
| 161 | rewrites \<open>lex_preordering.lex_less_eq (\<le>) (<) = ((\<le>) :: 'a list \<Rightarrow> 'a list \<Rightarrow> bool)\<close> | |
| 162 | and \<open>lex_preordering.lex_less (\<le>) (<) = ((<) :: 'a list \<Rightarrow> 'a list \<Rightarrow> bool)\<close> | |
| 163 | proof - | |
| 164 | interpret lex_ordering \<open>(\<le>) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close> \<open>(<) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close> .. | |
| 165 | show \<open>lex_ordering ((\<le>) :: 'a \<Rightarrow> 'a \<Rightarrow> bool) (<)\<close> | |
| 166 | by (fact lex_ordering_axioms) | |
| 167 | show \<open>lex_preordering.lex_less_eq (\<le>) (<) = (\<le>)\<close> | |
| 168 | by (simp add: less_eq_list_def) | |
| 169 | show \<open>lex_preordering.lex_less (\<le>) (<) = (<)\<close> | |
| 170 | by (simp add: less_list_def) | |
| 171 | qed | |
| 172 | ||
| 173 | instance list :: (order) order | |
| 77955 
c4677a6aae2c
more standard name bindings (amending 5bf71b4da706): avoid odd full_name like "Orderings.class.Orderings.preorder.of_class.intro" with many redundant name space accesses;
 wenzelm parents: 
75244diff
changeset | 174 | by (rule order.intro_of_class, rule ordering_orderI, fact lex.ordering) | 
| 73794 | 175 | |
| 176 | export_code \<open>(\<le>) :: _ list \<Rightarrow> _ list \<Rightarrow> bool\<close> \<open>(<) :: _ list \<Rightarrow> _ list \<Rightarrow> bool\<close> in Haskell | |
| 177 | ||
| 178 | ||
| 179 | subsection \<open>Non-canonical instance\<close> | |
| 180 | ||
| 181 | context comm_monoid_mult | |
| 182 | begin | |
| 183 | ||
| 184 | definition dvd_strict :: \<open>'a \<Rightarrow> 'a \<Rightarrow> bool\<close> | |
| 185 | where \<open>dvd_strict a b \<longleftrightarrow> a dvd b \<and> \<not> b dvd a\<close> | |
| 186 | ||
| 187 | end | |
| 188 | ||
| 189 | global_interpretation dvd: lex_preordering \<open>(dvd) :: 'a::comm_monoid_mult \<Rightarrow> 'a \<Rightarrow> bool\<close> dvd_strict | |
| 190 | defines lex_dvd = dvd.lex_less_eq | |
| 191 | and lex_dvd_strict = dvd.lex_less | |
| 75244 | 192 | by unfold_locales (auto simp add: dvd_strict_def) | 
| 73794 | 193 | |
| 194 | global_interpretation lex_dvd: preordering lex_dvd lex_dvd_strict | |
| 195 | by (fact dvd.preordering) | |
| 196 | ||
| 197 | definition \<open>example = lex_dvd [(4::int), - 7, 8] [- 8, 13, 5]\<close> | |
| 198 | ||
| 199 | export_code example in Haskell | |
| 200 | ||
| 201 | value example | |
| 202 | ||
| 203 | end |