src/HOL/Library/Sublist.thy
author haftmann
Thu, 25 Jul 2013 08:57:16 +0200
changeset 53866 412c9e0381a1
parent 51531 ed6b40d15d1c
child 54152 a1119cf551e8
permissions -rw-r--r--
factored syntactic type classes for bot and top (by Alessandro Coglio)
Christian@50102
     1
(*  Title:      HOL/Library/Sublist.thy
wenzelm@10330
     2
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
Christian@50102
     3
    Author:     Christian Sternagel, JAIST
wenzelm@10330
     4
*)
wenzelm@10330
     5
Christian@51531
     6
header {* List prefixes, suffixes, and homeomorphic embedding *}
wenzelm@10330
     7
Christian@50102
     8
theory Sublist
Christian@50102
     9
imports Main
nipkow@15131
    10
begin
wenzelm@10330
    11
wenzelm@10330
    12
subsection {* Prefix order on lists *}
wenzelm@10330
    13
Christian@51531
    14
definition prefixeq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
wenzelm@50122
    15
  where "prefixeq xs ys \<longleftrightarrow> (\<exists>zs. ys = xs @ zs)"
wenzelm@10330
    16
Christian@51531
    17
definition prefix :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
wenzelm@50122
    18
  where "prefix xs ys \<longleftrightarrow> prefixeq xs ys \<and> xs \<noteq> ys"
wenzelm@10330
    19
Christian@50102
    20
interpretation prefix_order: order prefixeq prefix
Christian@50102
    21
  by default (auto simp: prefixeq_def prefix_def)
haftmann@25764
    22
haftmann@53866
    23
interpretation prefix_bot: order_bot Nil prefixeq prefix
Christian@50102
    24
  by default (simp add: prefixeq_def)
haftmann@37449
    25
Christian@51531
    26
lemma prefixeqI [intro?]: "ys = xs @ zs \<Longrightarrow> prefixeq xs ys"
Christian@50102
    27
  unfolding prefixeq_def by blast
wenzelm@10389
    28
Christian@50102
    29
lemma prefixeqE [elim?]:
Christian@50102
    30
  assumes "prefixeq xs ys"
Christian@50102
    31
  obtains zs where "ys = xs @ zs"
Christian@50102
    32
  using assms unfolding prefixeq_def by blast
haftmann@25764
    33
Christian@51531
    34
lemma prefixI' [intro?]: "ys = xs @ z # zs \<Longrightarrow> prefix xs ys"
Christian@50102
    35
  unfolding prefix_def prefixeq_def by blast
Christian@50102
    36
Christian@50102
    37
lemma prefixE' [elim?]:
Christian@50102
    38
  assumes "prefix xs ys"
Christian@50102
    39
  obtains z zs where "ys = xs @ z # zs"
Christian@50102
    40
proof -
Christian@50102
    41
  from `prefix xs ys` obtain us where "ys = xs @ us" and "xs \<noteq> ys"
Christian@50102
    42
    unfolding prefix_def prefixeq_def by blast
Christian@50102
    43
  with that show ?thesis by (auto simp add: neq_Nil_conv)
Christian@50102
    44
qed
Christian@50102
    45
Christian@51531
    46
lemma prefixI [intro?]: "prefixeq xs ys \<Longrightarrow> xs \<noteq> ys \<Longrightarrow> prefix xs ys"
wenzelm@18730
    47
  unfolding prefix_def by blast
wenzelm@10389
    48
wenzelm@21305
    49
lemma prefixE [elim?]:
Christian@50102
    50
  fixes xs ys :: "'a list"
Christian@50102
    51
  assumes "prefix xs ys"
Christian@50102
    52
  obtains "prefixeq xs ys" and "xs \<noteq> ys"
wenzelm@23394
    53
  using assms unfolding prefix_def by blast
wenzelm@10389
    54
wenzelm@10389
    55
wenzelm@10389
    56
subsection {* Basic properties of prefixes *}
wenzelm@10389
    57
Christian@50102
    58
theorem Nil_prefixeq [iff]: "prefixeq [] xs"
Christian@50102
    59
  by (simp add: prefixeq_def)
wenzelm@10389
    60
Christian@50102
    61
theorem prefixeq_Nil [simp]: "(prefixeq xs []) = (xs = [])"
Christian@50102
    62
  by (induct xs) (simp_all add: prefixeq_def)
wenzelm@10389
    63
Christian@50102
    64
lemma prefixeq_snoc [simp]: "prefixeq xs (ys @ [y]) \<longleftrightarrow> xs = ys @ [y] \<or> prefixeq xs ys"
wenzelm@10330
    65
proof
Christian@50102
    66
  assume "prefixeq xs (ys @ [y])"
wenzelm@10389
    67
  then obtain zs where zs: "ys @ [y] = xs @ zs" ..
Christian@50102
    68
  show "xs = ys @ [y] \<or> prefixeq xs ys"
Christian@50102
    69
    by (metis append_Nil2 butlast_append butlast_snoc prefixeqI zs)
wenzelm@10389
    70
next
Christian@50102
    71
  assume "xs = ys @ [y] \<or> prefixeq xs ys"
Christian@50102
    72
  then show "prefixeq xs (ys @ [y])"
Christian@50102
    73
    by (metis prefix_order.eq_iff prefix_order.order_trans prefixeqI)
wenzelm@10330
    74
qed
wenzelm@10330
    75
Christian@50102
    76
lemma Cons_prefixeq_Cons [simp]: "prefixeq (x # xs) (y # ys) = (x = y \<and> prefixeq xs ys)"
Christian@50102
    77
  by (auto simp add: prefixeq_def)
wenzelm@10330
    78
Christian@50102
    79
lemma prefixeq_code [code]:
Christian@50102
    80
  "prefixeq [] xs \<longleftrightarrow> True"
Christian@50102
    81
  "prefixeq (x # xs) [] \<longleftrightarrow> False"
Christian@50102
    82
  "prefixeq (x # xs) (y # ys) \<longleftrightarrow> x = y \<and> prefixeq xs ys"
haftmann@37449
    83
  by simp_all
haftmann@37449
    84
Christian@50102
    85
lemma same_prefixeq_prefixeq [simp]: "prefixeq (xs @ ys) (xs @ zs) = prefixeq ys zs"
wenzelm@10389
    86
  by (induct xs) simp_all
wenzelm@10330
    87
Christian@50102
    88
lemma same_prefixeq_nil [iff]: "prefixeq (xs @ ys) xs = (ys = [])"
Christian@50102
    89
  by (metis append_Nil2 append_self_conv prefix_order.eq_iff prefixeqI)
nipkow@25665
    90
Christian@51531
    91
lemma prefixeq_prefixeq [simp]: "prefixeq xs ys \<Longrightarrow> prefixeq xs (ys @ zs)"
Christian@50102
    92
  by (metis prefix_order.le_less_trans prefixeqI prefixE prefixI)
nipkow@25665
    93
Christian@50102
    94
lemma append_prefixeqD: "prefixeq (xs @ ys) zs \<Longrightarrow> prefixeq xs zs"
Christian@50102
    95
  by (auto simp add: prefixeq_def)
nipkow@14300
    96
Christian@50102
    97
theorem prefixeq_Cons: "prefixeq xs (y # ys) = (xs = [] \<or> (\<exists>zs. xs = y # zs \<and> prefixeq zs ys))"
Christian@50102
    98
  by (cases xs) (auto simp add: prefixeq_def)
wenzelm@10330
    99
Christian@50102
   100
theorem prefixeq_append:
Christian@50102
   101
  "prefixeq xs (ys @ zs) = (prefixeq xs ys \<or> (\<exists>us. xs = ys @ us \<and> prefixeq us zs))"
wenzelm@10330
   102
  apply (induct zs rule: rev_induct)
wenzelm@10330
   103
   apply force
wenzelm@10330
   104
  apply (simp del: append_assoc add: append_assoc [symmetric])
nipkow@25564
   105
  apply (metis append_eq_appendI)
wenzelm@10330
   106
  done
wenzelm@10330
   107
Christian@50102
   108
lemma append_one_prefixeq:
Christian@51531
   109
  "prefixeq xs ys \<Longrightarrow> length xs < length ys \<Longrightarrow> prefixeq (xs @ [ys ! length xs]) ys"
Christian@50102
   110
  unfolding prefixeq_def
wenzelm@25692
   111
  by (metis Cons_eq_appendI append_eq_appendI append_eq_conv_conj
wenzelm@25692
   112
    eq_Nil_appendI nth_drop')
nipkow@25665
   113
Christian@51531
   114
theorem prefixeq_length_le: "prefixeq xs ys \<Longrightarrow> length xs \<le> length ys"
Christian@50102
   115
  by (auto simp add: prefixeq_def)
wenzelm@10330
   116
Christian@50102
   117
lemma prefixeq_same_cases:
Christian@50102
   118
  "prefixeq (xs\<^isub>1::'a list) ys \<Longrightarrow> prefixeq xs\<^isub>2 ys \<Longrightarrow> prefixeq xs\<^isub>1 xs\<^isub>2 \<or> prefixeq xs\<^isub>2 xs\<^isub>1"
Christian@50102
   119
  unfolding prefixeq_def by (metis append_eq_append_conv2)
nipkow@25665
   120
Christian@50102
   121
lemma set_mono_prefixeq: "prefixeq xs ys \<Longrightarrow> set xs \<subseteq> set ys"
Christian@50102
   122
  by (auto simp add: prefixeq_def)
nipkow@14300
   123
Christian@50102
   124
lemma take_is_prefixeq: "prefixeq (take n xs) xs"
Christian@50102
   125
  unfolding prefixeq_def by (metis append_take_drop_id)
nipkow@25665
   126
Christian@50102
   127
lemma map_prefixeqI: "prefixeq xs ys \<Longrightarrow> prefixeq (map f xs) (map f ys)"
Christian@50102
   128
  by (auto simp: prefixeq_def)
kleing@25322
   129
Christian@50102
   130
lemma prefixeq_length_less: "prefix xs ys \<Longrightarrow> length xs < length ys"
Christian@50102
   131
  by (auto simp: prefix_def prefixeq_def)
nipkow@25665
   132
Christian@50102
   133
lemma prefix_simps [simp, code]:
Christian@50102
   134
  "prefix xs [] \<longleftrightarrow> False"
Christian@50102
   135
  "prefix [] (x # xs) \<longleftrightarrow> True"
Christian@50102
   136
  "prefix (x # xs) (y # ys) \<longleftrightarrow> x = y \<and> prefix xs ys"
Christian@50102
   137
  by (simp_all add: prefix_def cong: conj_cong)
kleing@25299
   138
Christian@50102
   139
lemma take_prefix: "prefix xs ys \<Longrightarrow> prefix (take n xs) ys"
wenzelm@25692
   140
  apply (induct n arbitrary: xs ys)
wenzelm@25692
   141
   apply (case_tac ys, simp_all)[1]
Christian@50102
   142
  apply (metis prefix_order.less_trans prefixI take_is_prefixeq)
wenzelm@25692
   143
  done
kleing@25299
   144
Christian@50102
   145
lemma not_prefixeq_cases:
Christian@50102
   146
  assumes pfx: "\<not> prefixeq ps ls"
wenzelm@25356
   147
  obtains
wenzelm@25356
   148
    (c1) "ps \<noteq> []" and "ls = []"
Christian@50102
   149
  | (c2) a as x xs where "ps = a#as" and "ls = x#xs" and "x = a" and "\<not> prefixeq as xs"
wenzelm@25356
   150
  | (c3) a as x xs where "ps = a#as" and "ls = x#xs" and "x \<noteq> a"
kleing@25299
   151
proof (cases ps)
wenzelm@50122
   152
  case Nil
wenzelm@50122
   153
  then show ?thesis using pfx by simp
kleing@25299
   154
next
kleing@25299
   155
  case (Cons a as)
wenzelm@25692
   156
  note c = `ps = a#as`
kleing@25299
   157
  show ?thesis
kleing@25299
   158
  proof (cases ls)
Christian@50102
   159
    case Nil then show ?thesis by (metis append_Nil2 pfx c1 same_prefixeq_nil)
kleing@25299
   160
  next
kleing@25299
   161
    case (Cons x xs)
kleing@25299
   162
    show ?thesis
kleing@25299
   163
    proof (cases "x = a")
wenzelm@25355
   164
      case True
Christian@50102
   165
      have "\<not> prefixeq as xs" using pfx c Cons True by simp
wenzelm@25355
   166
      with c Cons True show ?thesis by (rule c2)
wenzelm@25355
   167
    next
wenzelm@25355
   168
      case False
wenzelm@25355
   169
      with c Cons show ?thesis by (rule c3)
kleing@25299
   170
    qed
kleing@25299
   171
  qed
kleing@25299
   172
qed
kleing@25299
   173
Christian@50102
   174
lemma not_prefixeq_induct [consumes 1, case_names Nil Neq Eq]:
Christian@50102
   175
  assumes np: "\<not> prefixeq ps ls"
wenzelm@25356
   176
    and base: "\<And>x xs. P (x#xs) []"
wenzelm@25356
   177
    and r1: "\<And>x xs y ys. x \<noteq> y \<Longrightarrow> P (x#xs) (y#ys)"
Christian@50102
   178
    and r2: "\<And>x xs y ys. \<lbrakk> x = y; \<not> prefixeq xs ys; P xs ys \<rbrakk> \<Longrightarrow> P (x#xs) (y#ys)"
wenzelm@25356
   179
  shows "P ps ls" using np
kleing@25299
   180
proof (induct ls arbitrary: ps)
wenzelm@25355
   181
  case Nil then show ?case
Christian@50102
   182
    by (auto simp: neq_Nil_conv elim!: not_prefixeq_cases intro!: base)
kleing@25299
   183
next
wenzelm@25355
   184
  case (Cons y ys)
Christian@50102
   185
  then have npfx: "\<not> prefixeq ps (y # ys)" by simp
wenzelm@25355
   186
  then obtain x xs where pv: "ps = x # xs"
Christian@50102
   187
    by (rule not_prefixeq_cases) auto
Christian@50102
   188
  show ?case by (metis Cons.hyps Cons_prefixeq_Cons npfx pv r1 r2)
kleing@25299
   189
qed
nipkow@14300
   190
wenzelm@25356
   191
wenzelm@10389
   192
subsection {* Parallel lists *}
wenzelm@10330
   193
Christian@51531
   194
definition parallel :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "\<parallel>" 50)
wenzelm@50122
   195
  where "(xs \<parallel> ys) = (\<not> prefixeq xs ys \<and> \<not> prefixeq ys xs)"
wenzelm@10330
   196
Christian@51531
   197
lemma parallelI [intro]: "\<not> prefixeq xs ys \<Longrightarrow> \<not> prefixeq ys xs \<Longrightarrow> xs \<parallel> ys"
wenzelm@25692
   198
  unfolding parallel_def by blast
wenzelm@10330
   199
wenzelm@10389
   200
lemma parallelE [elim]:
wenzelm@25692
   201
  assumes "xs \<parallel> ys"
Christian@50102
   202
  obtains "\<not> prefixeq xs ys \<and> \<not> prefixeq ys xs"
wenzelm@25692
   203
  using assms unfolding parallel_def by blast
wenzelm@10330
   204
Christian@50102
   205
theorem prefixeq_cases:
Christian@50102
   206
  obtains "prefixeq xs ys" | "prefix ys xs" | "xs \<parallel> ys"
Christian@50102
   207
  unfolding parallel_def prefix_def by blast
wenzelm@10330
   208
wenzelm@10389
   209
theorem parallel_decomp:
Christian@51531
   210
  "xs \<parallel> ys \<Longrightarrow> \<exists>as b bs c cs. b \<noteq> c \<and> xs = as @ b # bs \<and> ys = as @ c # cs"
wenzelm@10408
   211
proof (induct xs rule: rev_induct)
wenzelm@11987
   212
  case Nil
wenzelm@23254
   213
  then have False by auto
wenzelm@23254
   214
  then show ?case ..
wenzelm@10408
   215
next
wenzelm@11987
   216
  case (snoc x xs)
wenzelm@11987
   217
  show ?case
Christian@50102
   218
  proof (rule prefixeq_cases)
Christian@50102
   219
    assume le: "prefixeq xs ys"
wenzelm@10408
   220
    then obtain ys' where ys: "ys = xs @ ys'" ..
wenzelm@10408
   221
    show ?thesis
wenzelm@10408
   222
    proof (cases ys')
nipkow@25564
   223
      assume "ys' = []"
Christian@50102
   224
      then show ?thesis by (metis append_Nil2 parallelE prefixeqI snoc.prems ys)
wenzelm@10389
   225
    next
wenzelm@10408
   226
      fix c cs assume ys': "ys' = c # cs"
wenzelm@25692
   227
      then show ?thesis
Christian@50102
   228
        by (metis Cons_eq_appendI eq_Nil_appendI parallelE prefixeqI
Christian@50102
   229
          same_prefixeq_prefixeq snoc.prems ys)
wenzelm@10389
   230
    qed
wenzelm@10408
   231
  next
wenzelm@50122
   232
    assume "prefix ys xs"
wenzelm@50122
   233
    then have "prefixeq ys (xs @ [x])" by (simp add: prefix_def)
wenzelm@11987
   234
    with snoc have False by blast
wenzelm@23254
   235
    then show ?thesis ..
wenzelm@10408
   236
  next
wenzelm@10408
   237
    assume "xs \<parallel> ys"
wenzelm@11987
   238
    with snoc obtain as b bs c cs where neq: "(b::'a) \<noteq> c"
wenzelm@10408
   239
      and xs: "xs = as @ b # bs" and ys: "ys = as @ c # cs"
wenzelm@10408
   240
      by blast
wenzelm@10408
   241
    from xs have "xs @ [x] = as @ b # (bs @ [x])" by simp
wenzelm@10408
   242
    with neq ys show ?thesis by blast
wenzelm@10389
   243
  qed
wenzelm@10389
   244
qed
wenzelm@10330
   245
nipkow@25564
   246
lemma parallel_append: "a \<parallel> b \<Longrightarrow> a @ c \<parallel> b @ d"
wenzelm@25692
   247
  apply (rule parallelI)
wenzelm@25692
   248
    apply (erule parallelE, erule conjE,
Christian@50102
   249
      induct rule: not_prefixeq_induct, simp+)+
wenzelm@25692
   250
  done
kleing@25299
   251
wenzelm@25692
   252
lemma parallel_appendI: "xs \<parallel> ys \<Longrightarrow> x = xs @ xs' \<Longrightarrow> y = ys @ ys' \<Longrightarrow> x \<parallel> y"
wenzelm@25692
   253
  by (simp add: parallel_append)
kleing@25299
   254
wenzelm@25692
   255
lemma parallel_commute: "a \<parallel> b \<longleftrightarrow> b \<parallel> a"
wenzelm@25692
   256
  unfolding parallel_def by auto
oheimb@14538
   257
wenzelm@25356
   258
Christian@50102
   259
subsection {* Suffix order on lists *}
wenzelm@17201
   260
wenzelm@50122
   261
definition suffixeq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
wenzelm@50122
   262
  where "suffixeq xs ys = (\<exists>zs. ys = zs @ xs)"
oheimb@14538
   263
wenzelm@50122
   264
definition suffix :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
wenzelm@50122
   265
  where "suffix xs ys \<longleftrightarrow> (\<exists>us. ys = us @ xs \<and> us \<noteq> [])"
wenzelm@21305
   266
Christian@50102
   267
lemma suffix_imp_suffixeq:
Christian@50102
   268
  "suffix xs ys \<Longrightarrow> suffixeq xs ys"
Christian@50102
   269
  by (auto simp: suffixeq_def suffix_def)
wenzelm@21305
   270
Christian@51531
   271
lemma suffixeqI [intro?]: "ys = zs @ xs \<Longrightarrow> suffixeq xs ys"
Christian@50102
   272
  unfolding suffixeq_def by blast
oheimb@14538
   273
Christian@50102
   274
lemma suffixeqE [elim?]:
Christian@50102
   275
  assumes "suffixeq xs ys"
Christian@50102
   276
  obtains zs where "ys = zs @ xs"
Christian@50102
   277
  using assms unfolding suffixeq_def by blast
oheimb@14538
   278
Christian@50102
   279
lemma suffixeq_refl [iff]: "suffixeq xs xs"
Christian@50102
   280
  by (auto simp add: suffixeq_def)
Christian@50102
   281
lemma suffix_trans:
Christian@50102
   282
  "suffix xs ys \<Longrightarrow> suffix ys zs \<Longrightarrow> suffix xs zs"
Christian@50102
   283
  by (auto simp: suffix_def)
Christian@50102
   284
lemma suffixeq_trans: "\<lbrakk>suffixeq xs ys; suffixeq ys zs\<rbrakk> \<Longrightarrow> suffixeq xs zs"
Christian@50102
   285
  by (auto simp add: suffixeq_def)
Christian@50102
   286
lemma suffixeq_antisym: "\<lbrakk>suffixeq xs ys; suffixeq ys xs\<rbrakk> \<Longrightarrow> xs = ys"
Christian@50102
   287
  by (auto simp add: suffixeq_def)
oheimb@14538
   288
Christian@50102
   289
lemma suffixeq_tl [simp]: "suffixeq (tl xs) xs"
Christian@50102
   290
  by (induct xs) (auto simp: suffixeq_def)
oheimb@14538
   291
Christian@50102
   292
lemma suffix_tl [simp]: "xs \<noteq> [] \<Longrightarrow> suffix (tl xs) xs"
Christian@50102
   293
  by (induct xs) (auto simp: suffix_def)
Christian@50102
   294
Christian@50102
   295
lemma Nil_suffixeq [iff]: "suffixeq [] xs"
Christian@50102
   296
  by (simp add: suffixeq_def)
Christian@50102
   297
lemma suffixeq_Nil [simp]: "(suffixeq xs []) = (xs = [])"
Christian@50102
   298
  by (auto simp add: suffixeq_def)
Christian@50102
   299
wenzelm@50122
   300
lemma suffixeq_ConsI: "suffixeq xs ys \<Longrightarrow> suffixeq xs (y # ys)"
Christian@50102
   301
  by (auto simp add: suffixeq_def)
wenzelm@50122
   302
lemma suffixeq_ConsD: "suffixeq (x # xs) ys \<Longrightarrow> suffixeq xs ys"
Christian@50102
   303
  by (auto simp add: suffixeq_def)
Christian@50102
   304
Christian@50102
   305
lemma suffixeq_appendI: "suffixeq xs ys \<Longrightarrow> suffixeq xs (zs @ ys)"
Christian@50102
   306
  by (auto simp add: suffixeq_def)
Christian@50102
   307
lemma suffixeq_appendD: "suffixeq (zs @ xs) ys \<Longrightarrow> suffixeq xs ys"
Christian@50102
   308
  by (auto simp add: suffixeq_def)
Christian@50102
   309
Christian@50102
   310
lemma suffix_set_subset:
Christian@50102
   311
  "suffix xs ys \<Longrightarrow> set xs \<subseteq> set ys" by (auto simp: suffix_def)
Christian@50102
   312
Christian@50102
   313
lemma suffixeq_set_subset:
Christian@50102
   314
  "suffixeq xs ys \<Longrightarrow> set xs \<subseteq> set ys" by (auto simp: suffixeq_def)
Christian@50102
   315
wenzelm@50122
   316
lemma suffixeq_ConsD2: "suffixeq (x # xs) (y # ys) \<Longrightarrow> suffixeq xs ys"
wenzelm@21305
   317
proof -
wenzelm@50122
   318
  assume "suffixeq (x # xs) (y # ys)"
wenzelm@50122
   319
  then obtain zs where "y # ys = zs @ x # xs" ..
Christian@50102
   320
  then show ?thesis
Christian@50102
   321
    by (induct zs) (auto intro!: suffixeq_appendI suffixeq_ConsI)
wenzelm@21305
   322
qed
oheimb@14538
   323
Christian@50102
   324
lemma suffixeq_to_prefixeq [code]: "suffixeq xs ys \<longleftrightarrow> prefixeq (rev xs) (rev ys)"
Christian@50102
   325
proof
Christian@50102
   326
  assume "suffixeq xs ys"
Christian@50102
   327
  then obtain zs where "ys = zs @ xs" ..
Christian@50102
   328
  then have "rev ys = rev xs @ rev zs" by simp
Christian@50102
   329
  then show "prefixeq (rev xs) (rev ys)" ..
Christian@50102
   330
next
Christian@50102
   331
  assume "prefixeq (rev xs) (rev ys)"
Christian@50102
   332
  then obtain zs where "rev ys = rev xs @ zs" ..
Christian@50102
   333
  then have "rev (rev ys) = rev zs @ rev (rev xs)" by simp
Christian@50102
   334
  then have "ys = rev zs @ xs" by simp
Christian@50102
   335
  then show "suffixeq xs ys" ..
wenzelm@21305
   336
qed
oheimb@14538
   337
Christian@50102
   338
lemma distinct_suffixeq: "distinct ys \<Longrightarrow> suffixeq xs ys \<Longrightarrow> distinct xs"
Christian@50102
   339
  by (clarsimp elim!: suffixeqE)
wenzelm@17201
   340
Christian@50102
   341
lemma suffixeq_map: "suffixeq xs ys \<Longrightarrow> suffixeq (map f xs) (map f ys)"
Christian@50102
   342
  by (auto elim!: suffixeqE intro: suffixeqI)
kleing@25299
   343
Christian@50102
   344
lemma suffixeq_drop: "suffixeq (drop n as) as"
Christian@50102
   345
  unfolding suffixeq_def
wenzelm@25692
   346
  apply (rule exI [where x = "take n as"])
wenzelm@25692
   347
  apply simp
wenzelm@25692
   348
  done
kleing@25299
   349
Christian@50102
   350
lemma suffixeq_take: "suffixeq xs ys \<Longrightarrow> ys = take (length ys - length xs) ys @ xs"
wenzelm@50122
   351
  by (auto elim!: suffixeqE)
kleing@25299
   352
wenzelm@50122
   353
lemma suffixeq_suffix_reflclp_conv: "suffixeq = suffix\<^sup>=\<^sup>="
Christian@50102
   354
proof (intro ext iffI)
Christian@50102
   355
  fix xs ys :: "'a list"
Christian@50102
   356
  assume "suffixeq xs ys"
Christian@50102
   357
  show "suffix\<^sup>=\<^sup>= xs ys"
Christian@50102
   358
  proof
Christian@50102
   359
    assume "xs \<noteq> ys"
wenzelm@50122
   360
    with `suffixeq xs ys` show "suffix xs ys"
wenzelm@50122
   361
      by (auto simp: suffixeq_def suffix_def)
Christian@50102
   362
  qed
Christian@50102
   363
next
Christian@50102
   364
  fix xs ys :: "'a list"
Christian@50102
   365
  assume "suffix\<^sup>=\<^sup>= xs ys"
wenzelm@50122
   366
  then show "suffixeq xs ys"
Christian@50102
   367
  proof
wenzelm@50122
   368
    assume "suffix xs ys" then show "suffixeq xs ys"
wenzelm@50122
   369
      by (rule suffix_imp_suffixeq)
Christian@50102
   370
  next
wenzelm@50122
   371
    assume "xs = ys" then show "suffixeq xs ys"
wenzelm@50122
   372
      by (auto simp: suffixeq_def)
Christian@50102
   373
  qed
Christian@50102
   374
qed
Christian@50102
   375
Christian@50102
   376
lemma parallelD1: "x \<parallel> y \<Longrightarrow> \<not> prefixeq x y"
wenzelm@25692
   377
  by blast
kleing@25299
   378
Christian@50102
   379
lemma parallelD2: "x \<parallel> y \<Longrightarrow> \<not> prefixeq y x"
wenzelm@25692
   380
  by blast
wenzelm@25355
   381
wenzelm@25355
   382
lemma parallel_Nil1 [simp]: "\<not> x \<parallel> []"
wenzelm@25692
   383
  unfolding parallel_def by simp
wenzelm@25355
   384
kleing@25299
   385
lemma parallel_Nil2 [simp]: "\<not> [] \<parallel> x"
wenzelm@25692
   386
  unfolding parallel_def by simp
kleing@25299
   387
nipkow@25564
   388
lemma Cons_parallelI1: "a \<noteq> b \<Longrightarrow> a # as \<parallel> b # bs"
wenzelm@25692
   389
  by auto
kleing@25299
   390
nipkow@25564
   391
lemma Cons_parallelI2: "\<lbrakk> a = b; as \<parallel> bs \<rbrakk> \<Longrightarrow> a # as \<parallel> b # bs"
Christian@50102
   392
  by (metis Cons_prefixeq_Cons parallelE parallelI)
nipkow@25665
   393
kleing@25299
   394
lemma not_equal_is_parallel:
kleing@25299
   395
  assumes neq: "xs \<noteq> ys"
wenzelm@25356
   396
    and len: "length xs = length ys"
wenzelm@25356
   397
  shows "xs \<parallel> ys"
kleing@25299
   398
  using len neq
wenzelm@25355
   399
proof (induct rule: list_induct2)
haftmann@26445
   400
  case Nil
wenzelm@25356
   401
  then show ?case by simp
kleing@25299
   402
next
haftmann@26445
   403
  case (Cons a as b bs)
wenzelm@25355
   404
  have ih: "as \<noteq> bs \<Longrightarrow> as \<parallel> bs" by fact
kleing@25299
   405
  show ?case
kleing@25299
   406
  proof (cases "a = b")
wenzelm@25355
   407
    case True
haftmann@26445
   408
    then have "as \<noteq> bs" using Cons by simp
wenzelm@25355
   409
    then show ?thesis by (rule Cons_parallelI2 [OF True ih])
kleing@25299
   410
  next
kleing@25299
   411
    case False
wenzelm@25355
   412
    then show ?thesis by (rule Cons_parallelI1)
kleing@25299
   413
  qed
kleing@25299
   414
qed
haftmann@22178
   415
wenzelm@50122
   416
lemma suffix_reflclp_conv: "suffix\<^sup>=\<^sup>= = suffixeq"
Christian@50102
   417
  by (intro ext) (auto simp: suffixeq_def suffix_def)
Christian@50102
   418
wenzelm@50122
   419
lemma suffix_lists: "suffix xs ys \<Longrightarrow> ys \<in> lists A \<Longrightarrow> xs \<in> lists A"
Christian@50102
   420
  unfolding suffix_def by auto
Christian@50102
   421
Christian@50102
   422
Christian@51531
   423
subsection {* Homeomorphic embedding on lists *}
Christian@50102
   424
Christian@51531
   425
inductive list_hembeq :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
Christian@50102
   426
  for P :: "('a \<Rightarrow> 'a \<Rightarrow> bool)"
Christian@50102
   427
where
Christian@51531
   428
  list_hembeq_Nil [intro, simp]: "list_hembeq P [] ys"
Christian@51531
   429
| list_hembeq_Cons [intro] : "list_hembeq P xs ys \<Longrightarrow> list_hembeq P xs (y#ys)"
Christian@51531
   430
| list_hembeq_Cons2 [intro]: "P\<^sup>=\<^sup>= x y \<Longrightarrow> list_hembeq P xs ys \<Longrightarrow> list_hembeq P (x#xs) (y#ys)"
Christian@50102
   431
Christian@51531
   432
lemma list_hembeq_Nil2 [simp]:
Christian@51531
   433
  assumes "list_hembeq P xs []" shows "xs = []"
Christian@51531
   434
  using assms by (cases rule: list_hembeq.cases) auto
Christian@50102
   435
Christian@51531
   436
lemma list_hembeq_refl [simp, intro!]:
Christian@51531
   437
  "list_hembeq P xs xs"
Christian@51531
   438
  by (induct xs) auto
Christian@51531
   439
Christian@51531
   440
lemma list_hembeq_Cons_Nil [simp]: "list_hembeq P (x#xs) [] = False"
Christian@50102
   441
proof -
Christian@51531
   442
  { assume "list_hembeq P (x#xs) []"
Christian@51531
   443
    from list_hembeq_Nil2 [OF this] have False by simp
Christian@50102
   444
  } moreover {
Christian@50102
   445
    assume False
Christian@51531
   446
    then have "list_hembeq P (x#xs) []" by simp
Christian@50102
   447
  } ultimately show ?thesis by blast
Christian@50102
   448
qed
Christian@50102
   449
Christian@51531
   450
lemma list_hembeq_append2 [intro]: "list_hembeq P xs ys \<Longrightarrow> list_hembeq P xs (zs @ ys)"
Christian@50102
   451
  by (induct zs) auto
Christian@50102
   452
Christian@51531
   453
lemma list_hembeq_prefix [intro]:
Christian@51531
   454
  assumes "list_hembeq P xs ys" shows "list_hembeq P xs (ys @ zs)"
Christian@50102
   455
  using assms
Christian@50102
   456
  by (induct arbitrary: zs) auto
Christian@50102
   457
Christian@51531
   458
lemma list_hembeq_ConsD:
Christian@51531
   459
  assumes "list_hembeq P (x#xs) ys"
Christian@51531
   460
  shows "\<exists>us v vs. ys = us @ v # vs \<and> P\<^sup>=\<^sup>= x v \<and> list_hembeq P xs vs"
Christian@50102
   461
using assms
wenzelm@50122
   462
proof (induct x \<equiv> "x # xs" ys arbitrary: x xs)
Christian@51531
   463
  case list_hembeq_Cons
wenzelm@50122
   464
  then show ?case by (metis append_Cons)
Christian@50102
   465
next
Christian@51531
   466
  case (list_hembeq_Cons2 x y xs ys)
wenzelm@50122
   467
  then show ?case by (cases xs) (auto, blast+)
Christian@50102
   468
qed
Christian@50102
   469
Christian@51531
   470
lemma list_hembeq_appendD:
Christian@51531
   471
  assumes "list_hembeq P (xs @ ys) zs"
Christian@51531
   472
  shows "\<exists>us vs. zs = us @ vs \<and> list_hembeq P xs us \<and> list_hembeq P ys vs"
Christian@50102
   473
using assms
Christian@50102
   474
proof (induction xs arbitrary: ys zs)
wenzelm@50122
   475
  case Nil then show ?case by auto
Christian@50102
   476
next
Christian@50102
   477
  case (Cons x xs)
Christian@50102
   478
  then obtain us v vs where "zs = us @ v # vs"
Christian@51531
   479
    and "P\<^sup>=\<^sup>= x v" and "list_hembeq P (xs @ ys) vs" by (auto dest: list_hembeq_ConsD)
Christian@51531
   480
  with Cons show ?case by (metis append_Cons append_assoc list_hembeq_Cons2 list_hembeq_append2)
Christian@50102
   481
qed
Christian@50102
   482
Christian@51531
   483
lemma list_hembeq_suffix:
Christian@51531
   484
  assumes "list_hembeq P xs ys" and "suffix ys zs"
Christian@51531
   485
  shows "list_hembeq P xs zs"
Christian@51531
   486
  using assms(2) and list_hembeq_append2 [OF assms(1)] by (auto simp: suffix_def)
Christian@50102
   487
Christian@51531
   488
lemma list_hembeq_suffixeq:
Christian@51531
   489
  assumes "list_hembeq P xs ys" and "suffixeq ys zs"
Christian@51531
   490
  shows "list_hembeq P xs zs"
Christian@51531
   491
  using assms and list_hembeq_suffix unfolding suffixeq_suffix_reflclp_conv by auto
Christian@50102
   492
Christian@51531
   493
lemma list_hembeq_length: "list_hembeq P xs ys \<Longrightarrow> length xs \<le> length ys"
Christian@51531
   494
  by (induct rule: list_hembeq.induct) auto
Christian@50102
   495
Christian@51531
   496
lemma list_hembeq_trans:
Christian@51531
   497
  assumes "\<And>x y z. \<lbrakk>x \<in> A; y \<in> A; z \<in> A; P x y; P y z\<rbrakk> \<Longrightarrow> P x z"
Christian@51531
   498
  shows "\<And>xs ys zs. \<lbrakk>xs \<in> lists A; ys \<in> lists A; zs \<in> lists A;
Christian@51531
   499
    list_hembeq P xs ys; list_hembeq P ys zs\<rbrakk> \<Longrightarrow> list_hembeq P xs zs"
Christian@51531
   500
proof -
Christian@50102
   501
  fix xs ys zs
Christian@51531
   502
  assume "list_hembeq P xs ys" and "list_hembeq P ys zs"
Christian@50102
   503
    and "xs \<in> lists A" and "ys \<in> lists A" and "zs \<in> lists A"
Christian@51531
   504
  then show "list_hembeq P xs zs"
Christian@50102
   505
  proof (induction arbitrary: zs)
Christian@51531
   506
    case list_hembeq_Nil show ?case by blast
Christian@50102
   507
  next
Christian@51531
   508
    case (list_hembeq_Cons xs ys y)
Christian@51531
   509
    from list_hembeq_ConsD [OF `list_hembeq P (y#ys) zs`] obtain us v vs
Christian@51531
   510
      where zs: "zs = us @ v # vs" and "P\<^sup>=\<^sup>= y v" and "list_hembeq P ys vs" by blast
Christian@51531
   511
    then have "list_hembeq P ys (v#vs)" by blast
Christian@51531
   512
    then have "list_hembeq P ys zs" unfolding zs by (rule list_hembeq_append2)
Christian@51531
   513
    from list_hembeq_Cons.IH [OF this] and list_hembeq_Cons.prems show ?case by simp
Christian@50102
   514
  next
Christian@51531
   515
    case (list_hembeq_Cons2 x y xs ys)
Christian@51531
   516
    from list_hembeq_ConsD [OF `list_hembeq P (y#ys) zs`] obtain us v vs
Christian@51531
   517
      where zs: "zs = us @ v # vs" and "P\<^sup>=\<^sup>= y v" and "list_hembeq P ys vs" by blast
Christian@51531
   518
    with list_hembeq_Cons2 have "list_hembeq P xs vs" by simp
Christian@51531
   519
    moreover have "P\<^sup>=\<^sup>= x v"
Christian@50102
   520
    proof -
Christian@50102
   521
      from zs and `zs \<in> lists A` have "v \<in> A" by auto
Christian@51531
   522
      moreover have "x \<in> A" and "y \<in> A" using list_hembeq_Cons2 by simp_all
Christian@51531
   523
      ultimately show ?thesis
Christian@51531
   524
        using `P\<^sup>=\<^sup>= x y` and `P\<^sup>=\<^sup>= y v` and assms
Christian@51531
   525
        by blast
Christian@50102
   526
    qed
Christian@51531
   527
    ultimately have "list_hembeq P (x#xs) (v#vs)" by blast
Christian@51531
   528
    then show ?case unfolding zs by (rule list_hembeq_append2)
Christian@50102
   529
  qed
Christian@50102
   530
qed
Christian@50102
   531
Christian@50102
   532
Christian@51531
   533
subsection {* Sublists (special case of homeomorphic embedding) *}
Christian@50102
   534
Christian@51531
   535
abbreviation sublisteq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
Christian@51531
   536
  where "sublisteq xs ys \<equiv> list_hembeq (op =) xs ys"
Christian@50102
   537
Christian@51531
   538
lemma sublisteq_Cons2: "sublisteq xs ys \<Longrightarrow> sublisteq (x#xs) (x#ys)" by auto
Christian@50102
   539
Christian@51531
   540
lemma sublisteq_same_length:
Christian@51531
   541
  assumes "sublisteq xs ys" and "length xs = length ys" shows "xs = ys"
Christian@51531
   542
  using assms by (induct) (auto dest: list_hembeq_length)
Christian@50102
   543
Christian@51531
   544
lemma not_sublisteq_length [simp]: "length ys < length xs \<Longrightarrow> \<not> sublisteq xs ys"
Christian@51531
   545
  by (metis list_hembeq_length linorder_not_less)
Christian@50102
   546
Christian@50102
   547
lemma [code]:
Christian@51531
   548
  "list_hembeq P [] ys \<longleftrightarrow> True"
Christian@51531
   549
  "list_hembeq P (x#xs) [] \<longleftrightarrow> False"
Christian@50102
   550
  by (simp_all)
Christian@50102
   551
Christian@51531
   552
lemma sublisteq_Cons': "sublisteq (x#xs) ys \<Longrightarrow> sublisteq xs ys"
Christian@51531
   553
  by (induct xs) (auto dest: list_hembeq_ConsD)
Christian@50102
   554
Christian@51531
   555
lemma sublisteq_Cons2':
Christian@51531
   556
  assumes "sublisteq (x#xs) (x#ys)" shows "sublisteq xs ys"
Christian@51531
   557
  using assms by (cases) (rule sublisteq_Cons')
Christian@50102
   558
Christian@51531
   559
lemma sublisteq_Cons2_neq:
Christian@51531
   560
  assumes "sublisteq (x#xs) (y#ys)"
Christian@51531
   561
  shows "x \<noteq> y \<Longrightarrow> sublisteq (x#xs) ys"
Christian@50102
   562
  using assms by (cases) auto
Christian@50102
   563
Christian@51531
   564
lemma sublisteq_Cons2_iff [simp, code]:
Christian@51531
   565
  "sublisteq (x#xs) (y#ys) = (if x = y then sublisteq xs ys else sublisteq (x#xs) ys)"
Christian@51531
   566
  by (metis list_hembeq_Cons sublisteq_Cons2 sublisteq_Cons2' sublisteq_Cons2_neq)
Christian@50102
   567
Christian@51531
   568
lemma sublisteq_append': "sublisteq (zs @ xs) (zs @ ys) \<longleftrightarrow> sublisteq xs ys"
Christian@50102
   569
  by (induct zs) simp_all
Christian@50102
   570
Christian@51531
   571
lemma sublisteq_refl [simp, intro!]: "sublisteq xs xs" by (induct xs) simp_all
Christian@50102
   572
Christian@51531
   573
lemma sublisteq_antisym:
Christian@51531
   574
  assumes "sublisteq xs ys" and "sublisteq ys xs"
Christian@50102
   575
  shows "xs = ys"
Christian@50102
   576
using assms
Christian@50102
   577
proof (induct)
Christian@51531
   578
  case list_hembeq_Nil
Christian@51531
   579
  from list_hembeq_Nil2 [OF this] show ?case by simp
Christian@50102
   580
next
Christian@51531
   581
  case list_hembeq_Cons2
wenzelm@50122
   582
  then show ?case by simp
Christian@50102
   583
next
Christian@51531
   584
  case list_hembeq_Cons
wenzelm@50122
   585
  then show ?case
Christian@51531
   586
    by (metis sublisteq_Cons' list_hembeq_length Suc_length_conv Suc_n_not_le_n)
Christian@50102
   587
qed
Christian@50102
   588
Christian@51531
   589
lemma sublisteq_trans: "sublisteq xs ys \<Longrightarrow> sublisteq ys zs \<Longrightarrow> sublisteq xs zs"
Christian@51531
   590
  by (rule list_hembeq_trans [of UNIV "op ="]) auto
Christian@50102
   591
Christian@51531
   592
lemma sublisteq_append_le_same_iff: "sublisteq (xs @ ys) ys \<longleftrightarrow> xs = []"
Christian@51531
   593
  by (auto dest: list_hembeq_length)
Christian@50102
   594
Christian@51531
   595
lemma list_hembeq_append_mono:
Christian@51531
   596
  "\<lbrakk> list_hembeq P xs xs'; list_hembeq P ys ys' \<rbrakk> \<Longrightarrow> list_hembeq P (xs@ys) (xs'@ys')"
Christian@51531
   597
  apply (induct rule: list_hembeq.induct)
Christian@51531
   598
    apply (metis eq_Nil_appendI list_hembeq_append2)
Christian@51531
   599
   apply (metis append_Cons list_hembeq_Cons)
Christian@51531
   600
  apply (metis append_Cons list_hembeq_Cons2)
wenzelm@50122
   601
  done
Christian@50102
   602
Christian@50102
   603
Christian@50102
   604
subsection {* Appending elements *}
Christian@50102
   605
Christian@51531
   606
lemma sublisteq_append [simp]:
Christian@51531
   607
  "sublisteq (xs @ zs) (ys @ zs) \<longleftrightarrow> sublisteq xs ys" (is "?l = ?r")
Christian@50102
   608
proof
Christian@51531
   609
  { fix xs' ys' xs ys zs :: "'a list" assume "sublisteq xs' ys'"
Christian@51531
   610
    then have "xs' = xs @ zs & ys' = ys @ zs \<longrightarrow> sublisteq xs ys"
Christian@50102
   611
    proof (induct arbitrary: xs ys zs)
Christian@51531
   612
      case list_hembeq_Nil show ?case by simp
Christian@50102
   613
    next
Christian@51531
   614
      case (list_hembeq_Cons xs' ys' x)
Christian@51531
   615
      { assume "ys=[]" then have ?case using list_hembeq_Cons(1) by auto }
Christian@50102
   616
      moreover
Christian@50102
   617
      { fix us assume "ys = x#us"
Christian@51531
   618
        then have ?case using list_hembeq_Cons(2) by(simp add: list_hembeq.list_hembeq_Cons) }
Christian@50102
   619
      ultimately show ?case by (auto simp:Cons_eq_append_conv)
Christian@50102
   620
    next
Christian@51531
   621
      case (list_hembeq_Cons2 x y xs' ys')
Christian@51531
   622
      { assume "xs=[]" then have ?case using list_hembeq_Cons2(1) by auto }
Christian@50102
   623
      moreover
Christian@51531
   624
      { fix us vs assume "xs=x#us" "ys=x#vs" then have ?case using list_hembeq_Cons2 by auto}
Christian@50102
   625
      moreover
Christian@51531
   626
      { fix us assume "xs=x#us" "ys=[]" then have ?case using list_hembeq_Cons2(2) by bestsimp }
Christian@51531
   627
      ultimately show ?case using `op =\<^sup>=\<^sup>= x y` by (auto simp: Cons_eq_append_conv)
Christian@50102
   628
    qed }
Christian@50102
   629
  moreover assume ?l
Christian@50102
   630
  ultimately show ?r by blast
Christian@50102
   631
next
Christian@51531
   632
  assume ?r then show ?l by (metis list_hembeq_append_mono sublisteq_refl)
Christian@50102
   633
qed
Christian@50102
   634
Christian@51531
   635
lemma sublisteq_drop_many: "sublisteq xs ys \<Longrightarrow> sublisteq xs (zs @ ys)"
Christian@50102
   636
  by (induct zs) auto
Christian@50102
   637
Christian@51531
   638
lemma sublisteq_rev_drop_many: "sublisteq xs ys \<Longrightarrow> sublisteq xs (ys @ zs)"
Christian@51531
   639
  by (metis append_Nil2 list_hembeq_Nil list_hembeq_append_mono)
Christian@50102
   640
Christian@50102
   641
Christian@50102
   642
subsection {* Relation to standard list operations *}
Christian@50102
   643
Christian@51531
   644
lemma sublisteq_map:
Christian@51531
   645
  assumes "sublisteq xs ys" shows "sublisteq (map f xs) (map f ys)"
Christian@50102
   646
  using assms by (induct) auto
Christian@50102
   647
Christian@51531
   648
lemma sublisteq_filter_left [simp]: "sublisteq (filter P xs) xs"
Christian@50102
   649
  by (induct xs) auto
Christian@50102
   650
Christian@51531
   651
lemma sublisteq_filter [simp]:
Christian@51531
   652
  assumes "sublisteq xs ys" shows "sublisteq (filter P xs) (filter P ys)"
Christian@50102
   653
  using assms by (induct) auto
Christian@50102
   654
Christian@51531
   655
lemma "sublisteq xs ys \<longleftrightarrow> (\<exists>N. xs = sublist ys N)" (is "?L = ?R")
Christian@50102
   656
proof
Christian@50102
   657
  assume ?L
wenzelm@50122
   658
  then show ?R
Christian@50102
   659
  proof (induct)
Christian@51531
   660
    case list_hembeq_Nil show ?case by (metis sublist_empty)
Christian@50102
   661
  next
Christian@51531
   662
    case (list_hembeq_Cons xs ys x)
Christian@50102
   663
    then obtain N where "xs = sublist ys N" by blast
wenzelm@50122
   664
    then have "xs = sublist (x#ys) (Suc ` N)"
Christian@50102
   665
      by (clarsimp simp add:sublist_Cons inj_image_mem_iff)
wenzelm@50122
   666
    then show ?case by blast
Christian@50102
   667
  next
Christian@51531
   668
    case (list_hembeq_Cons2 x y xs ys)
Christian@50102
   669
    then obtain N where "xs = sublist ys N" by blast
wenzelm@50122
   670
    then have "x#xs = sublist (x#ys) (insert 0 (Suc ` N))"
Christian@50102
   671
      by (clarsimp simp add:sublist_Cons inj_image_mem_iff)
Christian@51531
   672
    moreover from list_hembeq_Cons2 have "x = y" by simp
Christian@51531
   673
    ultimately show ?case by blast
Christian@50102
   674
  qed
Christian@50102
   675
next
Christian@50102
   676
  assume ?R
Christian@50102
   677
  then obtain N where "xs = sublist ys N" ..
Christian@51531
   678
  moreover have "sublisteq (sublist ys N) ys"
wenzelm@50122
   679
  proof (induct ys arbitrary: N)
Christian@50102
   680
    case Nil show ?case by simp
Christian@50102
   681
  next
wenzelm@50122
   682
    case Cons then show ?case by (auto simp: sublist_Cons)
Christian@50102
   683
  qed
Christian@50102
   684
  ultimately show ?L by simp
Christian@50102
   685
qed
Christian@50102
   686
wenzelm@10330
   687
end