| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 18 Oct 2023 20:12:07 +0200 | |
| changeset 78843 | fc3ba0a1c82f | 
| parent 67613 | ce654b0e6d69 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32235diff
changeset | 1 | (* Title: HOL/UNITY/ListOrder.thy | 
| 6708 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1998 University of Cambridge | |
| 4 | ||
| 13798 | 5 | Lists are partially ordered by Charpentier's Generalized Prefix Relation | 
| 6 | (xs,ys) : genPrefix(r) | |
| 7 | if ys = xs' @ zs where length xs = length xs' | |
| 8 | and corresponding elements of xs, xs' are pairwise related by r | |
| 9 | ||
| 10 | Also overloads <= and < for lists! | |
| 6708 | 11 | *) | 
| 12 | ||
| 63146 | 13 | section \<open>The Prefix Ordering on Lists\<close> | 
| 13798 | 14 | |
| 27682 | 15 | theory ListOrder | 
| 16 | imports Main | |
| 17 | begin | |
| 13798 | 18 | |
| 23767 | 19 | inductive_set | 
| 13798 | 20 |   genPrefix :: "('a * 'a)set => ('a list * 'a list)set"
 | 
| 23767 | 21 |   for r :: "('a * 'a)set"
 | 
| 22 | where | |
| 67613 | 23 | Nil: "([],[]) \<in> genPrefix(r)" | 
| 13798 | 24 | |
| 67613 | 25 | | prepend: "[| (xs,ys) \<in> genPrefix(r); (x,y) \<in> r |] ==> | 
| 26 | (x#xs, y#ys) \<in> genPrefix(r)" | |
| 13798 | 27 | |
| 67613 | 28 | | append: "(xs,ys) \<in> genPrefix(r) ==> (xs, ys@zs) \<in> genPrefix(r)" | 
| 13798 | 29 | |
| 27682 | 30 | instantiation list :: (type) ord | 
| 31 | begin | |
| 13798 | 32 | |
| 27682 | 33 | definition | 
| 67613 | 34 | prefix_def: "xs <= zs \<longleftrightarrow> (xs, zs) \<in> genPrefix Id" | 
| 13798 | 35 | |
| 27682 | 36 | definition | 
| 37 | strict_prefix_def: "xs < zs \<longleftrightarrow> xs \<le> zs \<and> \<not> zs \<le> (xs :: 'a list)" | |
| 38 | ||
| 39 | instance .. | |
| 13798 | 40 | |
| 41 | (*Constants for the <= and >= relations, used below in translations*) | |
| 27682 | 42 | |
| 43 | end | |
| 44 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 45 | definition Le :: "(nat*nat) set" where | 
| 13798 | 46 |     "Le == {(x,y). x <= y}"
 | 
| 47 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 48 | definition Ge :: "(nat*nat) set" where | 
| 13798 | 49 |     "Ge == {(x,y). y <= x}"
 | 
| 50 | ||
| 23767 | 51 | abbreviation | 
| 52 | pfixLe :: "[nat list, nat list] => bool" (infixl "pfixLe" 50) where | |
| 67613 | 53 | "xs pfixLe ys == (xs,ys) \<in> genPrefix Le" | 
| 13798 | 54 | |
| 23767 | 55 | abbreviation | 
| 56 | pfixGe :: "[nat list, nat list] => bool" (infixl "pfixGe" 50) where | |
| 67613 | 57 | "xs pfixGe ys == (xs,ys) \<in> genPrefix Ge" | 
| 13798 | 58 | |
| 59 | ||
| 63146 | 60 | subsection\<open>preliminary lemmas\<close> | 
| 13798 | 61 | |
| 67613 | 62 | lemma Nil_genPrefix [iff]: "([], xs) \<in> genPrefix r" | 
| 13798 | 63 | by (cut_tac genPrefix.Nil [THEN genPrefix.append], auto) | 
| 64 | ||
| 67613 | 65 | lemma genPrefix_length_le: "(xs,ys) \<in> genPrefix r \<Longrightarrow> length xs <= length ys" | 
| 13798 | 66 | by (erule genPrefix.induct, auto) | 
| 67 | ||
| 68 | lemma cdlemma: | |
| 67613 | 69 | "[| (xs', ys') \<in> genPrefix r |] | 
| 70 | ==> (\<forall>x xs. xs' = x#xs \<longrightarrow> (\<exists>y ys. ys' = y#ys & (x,y) \<in> r & (xs, ys) \<in> genPrefix r))" | |
| 13798 | 71 | apply (erule genPrefix.induct, blast, blast) | 
| 72 | apply (force intro: genPrefix.append) | |
| 73 | done | |
| 74 | ||
| 75 | (*As usual converting it to an elimination rule is tiresome*) | |
| 76 | lemma cons_genPrefixE [elim!]: | |
| 67613 | 77 | "[| (x#xs, zs) \<in> genPrefix r; | 
| 78 | !!y ys. [| zs = y#ys; (x,y) \<in> r; (xs, ys) \<in> genPrefix r |] ==> P | |
| 13798 | 79 | |] ==> P" | 
| 80 | by (drule cdlemma, simp, blast) | |
| 81 | ||
| 82 | lemma Cons_genPrefix_Cons [iff]: | |
| 67613 | 83 | "((x#xs,y#ys) \<in> genPrefix r) = ((x,y) \<in> r \<and> (xs,ys) \<in> genPrefix r)" | 
| 13798 | 84 | by (blast intro: genPrefix.prepend) | 
| 85 | ||
| 86 | ||
| 63146 | 87 | subsection\<open>genPrefix is a partial order\<close> | 
| 13798 | 88 | |
| 30198 | 89 | lemma refl_genPrefix: "refl r ==> refl (genPrefix r)" | 
| 90 | apply (unfold refl_on_def, auto) | |
| 13798 | 91 | apply (induct_tac "x") | 
| 92 | prefer 2 apply (blast intro: genPrefix.prepend) | |
| 93 | apply (blast intro: genPrefix.Nil) | |
| 94 | done | |
| 95 | ||
| 67613 | 96 | lemma genPrefix_refl [simp]: "refl r \<Longrightarrow> (l,l) \<in> genPrefix r" | 
| 30198 | 97 | by (erule refl_onD [OF refl_genPrefix UNIV_I]) | 
| 13798 | 98 | |
| 99 | lemma genPrefix_mono: "r<=s ==> genPrefix r <= genPrefix s" | |
| 100 | apply clarify | |
| 101 | apply (erule genPrefix.induct) | |
| 102 | apply (auto intro: genPrefix.append) | |
| 103 | done | |
| 104 | ||
| 105 | ||
| 106 | (** Transitivity **) | |
| 107 | ||
| 108 | (*A lemma for proving genPrefix_trans_O*) | |
| 45477 | 109 | lemma append_genPrefix: | 
| 67613 | 110 | "(xs @ ys, zs) \<in> genPrefix r \<Longrightarrow> (xs, zs) \<in> genPrefix r" | 
| 45477 | 111 | by (induct xs arbitrary: zs) auto | 
| 13798 | 112 | |
| 113 | (*Lemma proving transitivity and more*) | |
| 45477 | 114 | lemma genPrefix_trans_O: | 
| 67613 | 115 | assumes "(x, y) \<in> genPrefix r" | 
| 116 | shows "\<And>z. (y, z) \<in> genPrefix s \<Longrightarrow> (x, z) \<in> genPrefix (r O s)" | |
| 45477 | 117 | apply (atomize (full)) | 
| 118 | using assms | |
| 119 | apply induct | |
| 120 | apply blast | |
| 121 | apply (blast intro: genPrefix.prepend) | |
| 122 | apply (blast dest: append_genPrefix) | |
| 123 | done | |
| 13798 | 124 | |
| 45477 | 125 | lemma genPrefix_trans: | 
| 67613 | 126 | "(x, y) \<in> genPrefix r \<Longrightarrow> (y, z) \<in> genPrefix r \<Longrightarrow> trans r | 
| 127 | \<Longrightarrow> (x, z) \<in> genPrefix r" | |
| 45477 | 128 | apply (rule trans_O_subset [THEN genPrefix_mono, THEN subsetD]) | 
| 129 | apply assumption | |
| 130 | apply (blast intro: genPrefix_trans_O) | |
| 131 | done | |
| 13798 | 132 | |
| 45477 | 133 | lemma prefix_genPrefix_trans: | 
| 67613 | 134 | "[| x<=y; (y,z) \<in> genPrefix r |] ==> (x, z) \<in> genPrefix r" | 
| 13798 | 135 | apply (unfold prefix_def) | 
| 32235 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 krauss parents: 
30198diff
changeset | 136 | apply (drule genPrefix_trans_O, assumption) | 
| 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 krauss parents: 
30198diff
changeset | 137 | apply simp | 
| 13798 | 138 | done | 
| 139 | ||
| 45477 | 140 | lemma genPrefix_prefix_trans: | 
| 67613 | 141 | "[| (x,y) \<in> genPrefix r; y<=z |] ==> (x,z) \<in> genPrefix r" | 
| 13798 | 142 | apply (unfold prefix_def) | 
| 32235 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 krauss parents: 
30198diff
changeset | 143 | apply (drule genPrefix_trans_O, assumption) | 
| 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 krauss parents: 
30198diff
changeset | 144 | apply simp | 
| 13798 | 145 | done | 
| 146 | ||
| 147 | lemma trans_genPrefix: "trans r ==> trans (genPrefix r)" | |
| 148 | by (blast intro: transI genPrefix_trans) | |
| 149 | ||
| 150 | ||
| 151 | (** Antisymmetry **) | |
| 152 | ||
| 45477 | 153 | lemma genPrefix_antisym: | 
| 67613 | 154 | assumes 1: "(xs, ys) \<in> genPrefix r" | 
| 45477 | 155 | and 2: "antisym r" | 
| 67613 | 156 | and 3: "(ys, xs) \<in> genPrefix r" | 
| 45477 | 157 | shows "xs = ys" | 
| 158 | using 1 3 | |
| 159 | proof induct | |
| 160 | case Nil | |
| 161 | then show ?case by blast | |
| 162 | next | |
| 163 | case prepend | |
| 164 | then show ?case using 2 by (simp add: antisym_def) | |
| 165 | next | |
| 166 | case (append xs ys zs) | |
| 167 | then show ?case | |
| 168 | apply - | |
| 169 | apply (subgoal_tac "length zs = 0", force) | |
| 170 | apply (drule genPrefix_length_le)+ | |
| 171 | apply (simp del: length_0_conv) | |
| 172 | done | |
| 173 | qed | |
| 13798 | 174 | |
| 175 | lemma antisym_genPrefix: "antisym r ==> antisym (genPrefix r)" | |
| 45477 | 176 | by (blast intro: antisymI genPrefix_antisym) | 
| 13798 | 177 | |
| 178 | ||
| 63146 | 179 | subsection\<open>recursion equations\<close> | 
| 13798 | 180 | |
| 67613 | 181 | lemma genPrefix_Nil [simp]: "((xs, []) \<in> genPrefix r) = (xs = [])" | 
| 46911 | 182 | by (induct xs) auto | 
| 13798 | 183 | |
| 184 | lemma same_genPrefix_genPrefix [simp]: | |
| 67613 | 185 | "refl r \<Longrightarrow> ((xs@ys, xs@zs) \<in> genPrefix r) = ((ys,zs) \<in> genPrefix r)" | 
| 46911 | 186 | by (induct xs) (simp_all add: refl_on_def) | 
| 13798 | 187 | |
| 188 | lemma genPrefix_Cons: | |
| 67613 | 189 | "((xs, y#ys) \<in> genPrefix r) = | 
| 190 | (xs=[] | (\<exists>z zs. xs=z#zs & (z,y) \<in> r & (zs,ys) \<in> genPrefix r))" | |
| 46911 | 191 | by (cases xs) auto | 
| 13798 | 192 | |
| 193 | lemma genPrefix_take_append: | |
| 67613 | 194 | "[| refl r; (xs,ys) \<in> genPrefix r |] | 
| 195 | ==> (xs@zs, take (length xs) ys @ zs) \<in> genPrefix r" | |
| 13798 | 196 | apply (erule genPrefix.induct) | 
| 197 | apply (frule_tac [3] genPrefix_length_le) | |
| 198 | apply (simp_all (no_asm_simp) add: diff_is_0_eq [THEN iffD2]) | |
| 199 | done | |
| 200 | ||
| 201 | lemma genPrefix_append_both: | |
| 67613 | 202 | "[| refl r; (xs,ys) \<in> genPrefix r; length xs = length ys |] | 
| 203 | ==> (xs@zs, ys @ zs) \<in> genPrefix r" | |
| 13798 | 204 | apply (drule genPrefix_take_append, assumption) | 
| 46577 | 205 | apply simp | 
| 13798 | 206 | done | 
| 207 | ||
| 208 | ||
| 209 | (*NOT suitable for rewriting since [y] has the form y#ys*) | |
| 210 | lemma append_cons_eq: "xs @ y # ys = (xs @ [y]) @ ys" | |
| 211 | by auto | |
| 212 | ||
| 213 | lemma aolemma: | |
| 67613 | 214 | "[| (xs,ys) \<in> genPrefix r; refl r |] | 
| 215 | ==> length xs < length ys \<longrightarrow> (xs @ [ys ! length xs], ys) \<in> genPrefix r" | |
| 13798 | 216 | apply (erule genPrefix.induct) | 
| 217 | apply blast | |
| 218 | apply simp | |
| 63146 | 219 | txt\<open>Append case is hardest\<close> | 
| 13798 | 220 | apply simp | 
| 221 | apply (frule genPrefix_length_le [THEN le_imp_less_or_eq]) | |
| 222 | apply (erule disjE) | |
| 223 | apply (simp_all (no_asm_simp) add: neq_Nil_conv nth_append) | |
| 224 | apply (blast intro: genPrefix.append, auto) | |
| 15481 | 225 | apply (subst append_cons_eq, fast intro: genPrefix_append_both genPrefix.append) | 
| 13798 | 226 | done | 
| 227 | ||
| 228 | lemma append_one_genPrefix: | |
| 67613 | 229 | "[| (xs,ys) \<in> genPrefix r; length xs < length ys; refl r |] | 
| 230 | ==> (xs @ [ys ! length xs], ys) \<in> genPrefix r" | |
| 13798 | 231 | by (blast intro: aolemma [THEN mp]) | 
| 232 | ||
| 233 | ||
| 234 | (** Proving the equivalence with Charpentier's definition **) | |
| 235 | ||
| 45477 | 236 | lemma genPrefix_imp_nth: | 
| 67613 | 237 | "i < length xs \<Longrightarrow> (xs, ys) \<in> genPrefix r \<Longrightarrow> (xs ! i, ys ! i) \<in> r" | 
| 45477 | 238 | apply (induct xs arbitrary: i ys) | 
| 239 | apply auto | |
| 240 | apply (case_tac i) | |
| 241 | apply auto | |
| 242 | done | |
| 13798 | 243 | |
| 45477 | 244 | lemma nth_imp_genPrefix: | 
| 245 | "length xs <= length ys \<Longrightarrow> | |
| 67613 | 246 | (\<forall>i. i < length xs \<longrightarrow> (xs ! i, ys ! i) \<in> r) \<Longrightarrow> | 
| 247 | (xs, ys) \<in> genPrefix r" | |
| 45477 | 248 | apply (induct xs arbitrary: ys) | 
| 249 | apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib) | |
| 250 | apply (case_tac ys) | |
| 251 | apply (force+) | |
| 252 | done | |
| 13798 | 253 | |
| 254 | lemma genPrefix_iff_nth: | |
| 67613 | 255 | "((xs,ys) \<in> genPrefix r) = | 
| 256 | (length xs <= length ys & (\<forall>i. i < length xs \<longrightarrow> (xs!i, ys!i) \<in> r))" | |
| 13798 | 257 | apply (blast intro: genPrefix_length_le genPrefix_imp_nth nth_imp_genPrefix) | 
| 258 | done | |
| 259 | ||
| 260 | ||
| 63146 | 261 | subsection\<open>The type of lists is partially ordered\<close> | 
| 13798 | 262 | |
| 30198 | 263 | declare refl_Id [iff] | 
| 13798 | 264 | antisym_Id [iff] | 
| 265 | trans_Id [iff] | |
| 266 | ||
| 267 | lemma prefix_refl [iff]: "xs <= (xs::'a list)" | |
| 268 | by (simp add: prefix_def) | |
| 269 | ||
| 270 | lemma prefix_trans: "!!xs::'a list. [| xs <= ys; ys <= zs |] ==> xs <= zs" | |
| 271 | apply (unfold prefix_def) | |
| 272 | apply (blast intro: genPrefix_trans) | |
| 273 | done | |
| 274 | ||
| 275 | lemma prefix_antisym: "!!xs::'a list. [| xs <= ys; ys <= xs |] ==> xs = ys" | |
| 276 | apply (unfold prefix_def) | |
| 277 | apply (blast intro: genPrefix_antisym) | |
| 278 | done | |
| 279 | ||
| 27682 | 280 | lemma prefix_less_le_not_le: "!!xs::'a list. (xs < zs) = (xs <= zs & \<not> zs \<le> xs)" | 
| 13798 | 281 | by (unfold strict_prefix_def, auto) | 
| 6708 | 282 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
6810diff
changeset | 283 | instance list :: (type) order | 
| 13798 | 284 | by (intro_classes, | 
| 285 | (assumption | rule prefix_refl prefix_trans prefix_antisym | |
| 27682 | 286 | prefix_less_le_not_le)+) | 
| 13798 | 287 | |
| 288 | (*Monotonicity of "set" operator WRT prefix*) | |
| 289 | lemma set_mono: "xs <= ys ==> set xs <= set ys" | |
| 290 | apply (unfold prefix_def) | |
| 291 | apply (erule genPrefix.induct, auto) | |
| 292 | done | |
| 293 | ||
| 294 | ||
| 295 | (** recursion equations **) | |
| 296 | ||
| 297 | lemma Nil_prefix [iff]: "[] <= xs" | |
| 46577 | 298 | by (simp add: prefix_def) | 
| 13798 | 299 | |
| 300 | lemma prefix_Nil [simp]: "(xs <= []) = (xs = [])" | |
| 46577 | 301 | by (simp add: prefix_def) | 
| 13798 | 302 | |
| 303 | lemma Cons_prefix_Cons [simp]: "(x#xs <= y#ys) = (x=y & xs<=ys)" | |
| 304 | by (simp add: prefix_def) | |
| 305 | ||
| 306 | lemma same_prefix_prefix [simp]: "(xs@ys <= xs@zs) = (ys <= zs)" | |
| 307 | by (simp add: prefix_def) | |
| 308 | ||
| 309 | lemma append_prefix [iff]: "(xs@ys <= xs) = (ys <= [])" | |
| 310 | by (insert same_prefix_prefix [of xs ys "[]"], simp) | |
| 311 | ||
| 312 | lemma prefix_appendI [simp]: "xs <= ys ==> xs <= ys@zs" | |
| 313 | apply (unfold prefix_def) | |
| 314 | apply (erule genPrefix.append) | |
| 315 | done | |
| 316 | ||
| 317 | lemma prefix_Cons: | |
| 67613 | 318 | "(xs <= y#ys) = (xs=[] | (\<exists>zs. xs=y#zs \<and> zs <= ys))" | 
| 13798 | 319 | by (simp add: prefix_def genPrefix_Cons) | 
| 320 | ||
| 321 | lemma append_one_prefix: | |
| 322 | "[| xs <= ys; length xs < length ys |] ==> xs @ [ys ! length xs] <= ys" | |
| 323 | apply (unfold prefix_def) | |
| 324 | apply (simp add: append_one_genPrefix) | |
| 325 | done | |
| 326 | ||
| 327 | lemma prefix_length_le: "xs <= ys ==> length xs <= length ys" | |
| 328 | apply (unfold prefix_def) | |
| 329 | apply (erule genPrefix_length_le) | |
| 330 | done | |
| 331 | ||
| 332 | lemma splemma: "xs<=ys ==> xs~=ys --> length xs < length ys" | |
| 333 | apply (unfold prefix_def) | |
| 334 | apply (erule genPrefix.induct, auto) | |
| 335 | done | |
| 336 | ||
| 337 | lemma strict_prefix_length_less: "xs < ys ==> length xs < length ys" | |
| 338 | apply (unfold strict_prefix_def) | |
| 339 | apply (blast intro: splemma [THEN mp]) | |
| 340 | done | |
| 341 | ||
| 342 | lemma mono_length: "mono length" | |
| 343 | by (blast intro: monoI prefix_length_le) | |
| 344 | ||
| 345 | (*Equivalence to the definition used in Lex/Prefix.thy*) | |
| 67613 | 346 | lemma prefix_iff: "(xs <= zs) = (\<exists>ys. zs = xs@ys)" | 
| 13798 | 347 | apply (unfold prefix_def) | 
| 348 | apply (auto simp add: genPrefix_iff_nth nth_append) | |
| 349 | apply (rule_tac x = "drop (length xs) zs" in exI) | |
| 350 | apply (rule nth_equalityI) | |
| 351 | apply (simp_all (no_asm_simp) add: nth_append) | |
| 352 | done | |
| 353 | ||
| 354 | lemma prefix_snoc [simp]: "(xs <= ys@[y]) = (xs = ys@[y] | xs <= ys)" | |
| 355 | apply (simp add: prefix_iff) | |
| 356 | apply (rule iffI) | |
| 357 | apply (erule exE) | |
| 358 | apply (rename_tac "zs") | |
| 359 | apply (rule_tac xs = zs in rev_exhaust) | |
| 360 | apply simp | |
| 361 | apply clarify | |
| 362 | apply (simp del: append_assoc add: append_assoc [symmetric], force) | |
| 363 | done | |
| 364 | ||
| 365 | lemma prefix_append_iff: | |
| 67613 | 366 | "(xs <= ys@zs) = (xs <= ys | (\<exists>us. xs = ys@us & us <= zs))" | 
| 13798 | 367 | apply (rule_tac xs = zs in rev_induct) | 
| 368 | apply force | |
| 369 | apply (simp del: append_assoc add: append_assoc [symmetric], force) | |
| 370 | done | |
| 371 | ||
| 372 | (*Although the prefix ordering is not linear, the prefixes of a list | |
| 373 | are linearly ordered.*) | |
| 45477 | 374 | lemma common_prefix_linear: | 
| 375 | fixes xs ys zs :: "'a list" | |
| 376 | shows "xs <= zs \<Longrightarrow> ys <= zs \<Longrightarrow> xs <= ys | ys <= xs" | |
| 377 | by (induct zs rule: rev_induct) auto | |
| 13798 | 378 | |
| 63146 | 379 | subsection\<open>pfixLe, pfixGe: properties inherited from the translations\<close> | 
| 13798 | 380 | |
| 381 | (** pfixLe **) | |
| 382 | ||
| 30198 | 383 | lemma refl_Le [iff]: "refl Le" | 
| 384 | by (unfold refl_on_def Le_def, auto) | |
| 13798 | 385 | |
| 386 | lemma antisym_Le [iff]: "antisym Le" | |
| 387 | by (unfold antisym_def Le_def, auto) | |
| 388 | ||
| 389 | lemma trans_Le [iff]: "trans Le" | |
| 390 | by (unfold trans_def Le_def, auto) | |
| 391 | ||
| 392 | lemma pfixLe_refl [iff]: "x pfixLe x" | |
| 393 | by simp | |
| 394 | ||
| 395 | lemma pfixLe_trans: "[| x pfixLe y; y pfixLe z |] ==> x pfixLe z" | |
| 396 | by (blast intro: genPrefix_trans) | |
| 397 | ||
| 398 | lemma pfixLe_antisym: "[| x pfixLe y; y pfixLe x |] ==> x = y" | |
| 399 | by (blast intro: genPrefix_antisym) | |
| 400 | ||
| 401 | lemma prefix_imp_pfixLe: "xs<=ys ==> xs pfixLe ys" | |
| 402 | apply (unfold prefix_def Le_def) | |
| 403 | apply (blast intro: genPrefix_mono [THEN [2] rev_subsetD]) | |
| 404 | done | |
| 405 | ||
| 30198 | 406 | lemma refl_Ge [iff]: "refl Ge" | 
| 407 | by (unfold refl_on_def Ge_def, auto) | |
| 13798 | 408 | |
| 409 | lemma antisym_Ge [iff]: "antisym Ge" | |
| 410 | by (unfold antisym_def Ge_def, auto) | |
| 411 | ||
| 412 | lemma trans_Ge [iff]: "trans Ge" | |
| 413 | by (unfold trans_def Ge_def, auto) | |
| 414 | ||
| 415 | lemma pfixGe_refl [iff]: "x pfixGe x" | |
| 416 | by simp | |
| 417 | ||
| 418 | lemma pfixGe_trans: "[| x pfixGe y; y pfixGe z |] ==> x pfixGe z" | |
| 419 | by (blast intro: genPrefix_trans) | |
| 420 | ||
| 421 | lemma pfixGe_antisym: "[| x pfixGe y; y pfixGe x |] ==> x = y" | |
| 422 | by (blast intro: genPrefix_antisym) | |
| 423 | ||
| 424 | lemma prefix_imp_pfixGe: "xs<=ys ==> xs pfixGe ys" | |
| 425 | apply (unfold prefix_def Ge_def) | |
| 426 | apply (blast intro: genPrefix_mono [THEN [2] rev_subsetD]) | |
| 427 | done | |
| 6708 | 428 | |
| 429 | end |