src/HOL/List_Prefix.thy
author haftmann
Sat, 25 Jan 2014 23:50:49 +0100
changeset 55147 bce3dbc11f95
parent 54538 ba7392b52a7c
permissions -rw-r--r--
prefer explicit code symbol type over ad-hoc name mangling
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54538
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/List_Prefix.thy
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     2
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     3
    Author:     Christian Sternagel, JAIST
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     4
*)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     5
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     6
header {* Parallel lists, list suffixes, and homeomorphic embedding *}
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     7
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     8
theory List_Prefix
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
     9
imports List
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    10
begin
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    11
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    12
subsection {* Prefix order on lists *}
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    13
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    14
definition prefixeq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    15
  where "prefixeq xs ys \<longleftrightarrow> (\<exists>zs. ys = xs @ zs)"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    16
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    17
definition prefix :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    18
  where "prefix xs ys \<longleftrightarrow> prefixeq xs ys \<and> xs \<noteq> ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    19
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    20
interpretation prefix_order: order prefixeq prefix
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    21
  by default (auto simp: prefixeq_def prefix_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    22
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    23
interpretation prefix_bot: order_bot Nil prefixeq prefix
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    24
  by default (simp add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    25
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    26
lemma prefixeqI [intro?]: "ys = xs @ zs \<Longrightarrow> prefixeq xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    27
  unfolding prefixeq_def by blast
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    28
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    29
lemma prefixeqE [elim?]:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    30
  assumes "prefixeq xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    31
  obtains zs where "ys = xs @ zs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    32
  using assms unfolding prefixeq_def by blast
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    33
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    34
lemma prefixI' [intro?]: "ys = xs @ z # zs \<Longrightarrow> prefix xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    35
  unfolding prefix_def prefixeq_def by blast
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    36
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    37
lemma prefixE' [elim?]:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    38
  assumes "prefix xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    39
  obtains z zs where "ys = xs @ z # zs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    40
proof -
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    41
  from `prefix xs ys` obtain us where "ys = xs @ us" and "xs \<noteq> ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    42
    unfolding prefix_def prefixeq_def by blast
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    43
  with that show ?thesis by (auto simp add: neq_Nil_conv)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    44
qed
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    45
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    46
lemma prefixI [intro?]: "prefixeq xs ys \<Longrightarrow> xs \<noteq> ys \<Longrightarrow> prefix xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    47
  unfolding prefix_def by blast
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    48
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    49
lemma prefixE [elim?]:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    50
  fixes xs ys :: "'a list"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    51
  assumes "prefix xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    52
  obtains "prefixeq xs ys" and "xs \<noteq> ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    53
  using assms unfolding prefix_def by blast
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    54
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    55
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    56
subsection {* Basic properties of prefixes *}
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    57
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    58
theorem Nil_prefixeq [iff]: "prefixeq [] xs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    59
  by (simp add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    60
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    61
theorem prefixeq_Nil [simp]: "(prefixeq xs []) = (xs = [])"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    62
  by (induct xs) (simp_all add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    63
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    64
lemma prefixeq_snoc [simp]: "prefixeq xs (ys @ [y]) \<longleftrightarrow> xs = ys @ [y] \<or> prefixeq xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    65
proof
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    66
  assume "prefixeq xs (ys @ [y])"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    67
  then obtain zs where zs: "ys @ [y] = xs @ zs" ..
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    68
  show "xs = ys @ [y] \<or> prefixeq xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    69
    by (metis append_Nil2 butlast_append butlast_snoc prefixeqI zs)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    70
next
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    71
  assume "xs = ys @ [y] \<or> prefixeq xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    72
  then show "prefixeq xs (ys @ [y])"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    73
    by (metis prefix_order.eq_iff prefix_order.order_trans prefixeqI)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    74
qed
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    75
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    76
lemma Cons_prefixeq_Cons [simp]: "prefixeq (x # xs) (y # ys) = (x = y \<and> prefixeq xs ys)"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    77
  by (auto simp add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    78
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    79
lemma prefixeq_code [code]:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    80
  "prefixeq [] xs \<longleftrightarrow> True"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    81
  "prefixeq (x # xs) [] \<longleftrightarrow> False"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    82
  "prefixeq (x # xs) (y # ys) \<longleftrightarrow> x = y \<and> prefixeq xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    83
  by simp_all
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    84
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    85
lemma same_prefixeq_prefixeq [simp]: "prefixeq (xs @ ys) (xs @ zs) = prefixeq ys zs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    86
  by (induct xs) simp_all
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    87
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    88
lemma same_prefixeq_nil [iff]: "prefixeq (xs @ ys) xs = (ys = [])"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    89
  by (metis append_Nil2 append_self_conv prefix_order.eq_iff prefixeqI)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    90
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    91
lemma prefixeq_prefixeq [simp]: "prefixeq xs ys \<Longrightarrow> prefixeq xs (ys @ zs)"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    92
  by (metis prefix_order.le_less_trans prefixeqI prefixE prefixI)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    93
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    94
lemma append_prefixeqD: "prefixeq (xs @ ys) zs \<Longrightarrow> prefixeq xs zs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    95
  by (auto simp add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    96
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    97
theorem prefixeq_Cons: "prefixeq xs (y # ys) = (xs = [] \<or> (\<exists>zs. xs = y # zs \<and> prefixeq zs ys))"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    98
  by (cases xs) (auto simp add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
    99
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   100
theorem prefixeq_append:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   101
  "prefixeq xs (ys @ zs) = (prefixeq xs ys \<or> (\<exists>us. xs = ys @ us \<and> prefixeq us zs))"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   102
  apply (induct zs rule: rev_induct)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   103
   apply force
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   104
  apply (simp del: append_assoc add: append_assoc [symmetric])
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   105
  apply (metis append_eq_appendI)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   106
  done
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   107
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   108
lemma append_one_prefixeq:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   109
  "prefixeq xs ys \<Longrightarrow> length xs < length ys \<Longrightarrow> prefixeq (xs @ [ys ! length xs]) ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   110
  proof (unfold prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   111
    assume a1: "\<exists>zs. ys = xs @ zs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   112
    then obtain sk :: "'a list" where sk: "ys = xs @ sk" by fastforce
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   113
    assume a2: "length xs < length ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   114
    have f1: "\<And>v. ([]\<Colon>'a list) @ v = v" using append_Nil2 by simp
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   115
    have "[] \<noteq> sk" using a1 a2 sk less_not_refl by force
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   116
    hence "\<exists>v. xs @ hd sk # v = ys" using sk by (metis hd_Cons_tl)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   117
    thus "\<exists>zs. ys = (xs @ [ys ! length xs]) @ zs" using f1 by fastforce
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   118
  qed
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   119
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   120
theorem prefixeq_length_le: "prefixeq xs ys \<Longrightarrow> length xs \<le> length ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   121
  by (auto simp add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   122
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   123
lemma prefixeq_same_cases:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   124
  "prefixeq (xs\<^sub>1::'a list) ys \<Longrightarrow> prefixeq xs\<^sub>2 ys \<Longrightarrow> prefixeq xs\<^sub>1 xs\<^sub>2 \<or> prefixeq xs\<^sub>2 xs\<^sub>1"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   125
  unfolding prefixeq_def by (force simp: append_eq_append_conv2)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   126
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   127
lemma set_mono_prefixeq: "prefixeq xs ys \<Longrightarrow> set xs \<subseteq> set ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   128
  by (auto simp add: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   129
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   130
lemma take_is_prefixeq: "prefixeq (take n xs) xs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   131
  unfolding prefixeq_def by (metis append_take_drop_id)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   132
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   133
lemma map_prefixeqI: "prefixeq xs ys \<Longrightarrow> prefixeq (map f xs) (map f ys)"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   134
  by (auto simp: prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   135
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   136
lemma prefixeq_length_less: "prefix xs ys \<Longrightarrow> length xs < length ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   137
  by (auto simp: prefix_def prefixeq_def)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   138
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   139
lemma prefix_simps [simp, code]:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   140
  "prefix xs [] \<longleftrightarrow> False"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   141
  "prefix [] (x # xs) \<longleftrightarrow> True"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   142
  "prefix (x # xs) (y # ys) \<longleftrightarrow> x = y \<and> prefix xs ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   143
  by (simp_all add: prefix_def cong: conj_cong)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   144
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   145
lemma take_prefix: "prefix xs ys \<Longrightarrow> prefix (take n xs) ys"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   146
  apply (induct n arbitrary: xs ys)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   147
   apply (case_tac ys, simp_all)[1]
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   148
  apply (metis prefix_order.less_trans prefixI take_is_prefixeq)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   149
  done
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   150
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   151
lemma not_prefixeq_cases:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   152
  assumes pfx: "\<not> prefixeq ps ls"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   153
  obtains
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   154
    (c1) "ps \<noteq> []" and "ls = []"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   155
  | (c2) a as x xs where "ps = a#as" and "ls = x#xs" and "x = a" and "\<not> prefixeq as xs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   156
  | (c3) a as x xs where "ps = a#as" and "ls = x#xs" and "x \<noteq> a"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   157
proof (cases ps)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   158
  case Nil
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   159
  then show ?thesis using pfx by simp
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   160
next
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   161
  case (Cons a as)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   162
  note c = `ps = a#as`
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   163
  show ?thesis
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   164
  proof (cases ls)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   165
    case Nil then show ?thesis by (metis append_Nil2 pfx c1 same_prefixeq_nil)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   166
  next
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   167
    case (Cons x xs)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   168
    show ?thesis
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   169
    proof (cases "x = a")
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   170
      case True
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   171
      have "\<not> prefixeq as xs" using pfx c Cons True by simp
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   172
      with c Cons True show ?thesis by (rule c2)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   173
    next
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   174
      case False
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   175
      with c Cons show ?thesis by (rule c3)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   176
    qed
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   177
  qed
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   178
qed
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   179
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   180
lemma not_prefixeq_induct [consumes 1, case_names Nil Neq Eq]:
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   181
  assumes np: "\<not> prefixeq ps ls"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   182
    and base: "\<And>x xs. P (x#xs) []"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   183
    and r1: "\<And>x xs y ys. x \<noteq> y \<Longrightarrow> P (x#xs) (y#ys)"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   184
    and r2: "\<And>x xs y ys. \<lbrakk> x = y; \<not> prefixeq xs ys; P xs ys \<rbrakk> \<Longrightarrow> P (x#xs) (y#ys)"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   185
  shows "P ps ls" using np
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   186
proof (induct ls arbitrary: ps)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   187
  case Nil then show ?case
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   188
    by (auto simp: neq_Nil_conv elim!: not_prefixeq_cases intro!: base)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   189
next
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   190
  case (Cons y ys)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   191
  then have npfx: "\<not> prefixeq ps (y # ys)" by simp
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   192
  then obtain x xs where pv: "ps = x # xs"
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   193
    by (rule not_prefixeq_cases) auto
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   194
  show ?case by (metis Cons.hyps Cons_prefixeq_Cons npfx pv r1 r2)
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   195
qed
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   196
ba7392b52a7c factor 'List_Prefix' out of 'Sublist' and move to 'Main' (needed for codatatypes)
blanchet
parents:
diff changeset
   197
end