src/HOL/Library/More_List.thy
author wenzelm
Sun, 21 Aug 2022 12:19:38 +0200
changeset 75942 603852abed8f
parent 75008 43b3d5318d72
child 81347 31f9e5ada550
permissions -rw-r--r--
clarified names: Browser_Info.Config vs. Browser_Info.Context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     1
(* Author: Andreas Lochbihler, ETH Zürich
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     2
   Author: Florian Haftmann, TU Muenchen  *)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     3
58881
b9556a055632 modernized header;
wenzelm
parents: 58437
diff changeset
     4
section \<open>Less common functions on lists\<close>
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     5
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     6
theory More_List
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     7
imports Main
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     8
begin
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
     9
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    10
definition strip_while :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    11
where
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    12
  "strip_while P = rev \<circ> dropWhile P \<circ> rev"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    13
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    14
lemma strip_while_rev [simp]:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    15
  "strip_while P (rev xs) = rev (dropWhile P xs)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    16
  by (simp add: strip_while_def)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    17
  
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    18
lemma strip_while_Nil [simp]:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    19
  "strip_while P [] = []"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    20
  by (simp add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    21
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    22
lemma strip_while_append [simp]:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    23
  "\<not> P x \<Longrightarrow> strip_while P (xs @ [x]) = xs @ [x]"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    24
  by (simp add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    25
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    26
lemma strip_while_append_rec [simp]:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    27
  "P x \<Longrightarrow> strip_while P (xs @ [x]) = strip_while P xs"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    28
  by (simp add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    29
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    30
lemma strip_while_Cons [simp]:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    31
  "\<not> P x \<Longrightarrow> strip_while P (x # xs) = x # strip_while P xs"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    32
  by (induct xs rule: rev_induct) (simp_all add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    33
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    34
lemma strip_while_eq_Nil [simp]:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    35
  "strip_while P xs = [] \<longleftrightarrow> (\<forall>x\<in>set xs. P x)"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    36
  by (simp add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    37
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    38
lemma strip_while_eq_Cons_rec:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    39
  "strip_while P (x # xs) = x # strip_while P xs \<longleftrightarrow> \<not> (P x \<and> (\<forall>x\<in>set xs. P x))"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    40
  by (induct xs rule: rev_induct) (simp_all add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    41
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    42
lemma split_strip_while_append:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    43
  fixes xs :: "'a list"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    44
  obtains ys zs :: "'a list"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    45
  where "strip_while P xs = ys" and "\<forall>x\<in>set zs. P x" and "xs = ys @ zs"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    46
proof (rule that)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    47
  show "strip_while P xs = strip_while P xs" ..
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    48
  show "\<forall>x\<in>set (rev (takeWhile P (rev xs))). P x" by (simp add: takeWhile_eq_all_conv [symmetric])
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    49
  have "rev xs = rev (strip_while P xs @ rev (takeWhile P (rev xs)))"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    50
    by (simp add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    51
  then show "xs = strip_while P xs @ rev (takeWhile P (rev xs))"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    52
    by (simp only: rev_is_rev_conv)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    53
qed
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    54
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    55
lemma strip_while_snoc [simp]:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    56
  "strip_while P (xs @ [x]) = (if P x then strip_while P xs else xs @ [x])"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    57
  by (simp add: strip_while_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    58
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    59
lemma strip_while_map:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    60
  "strip_while P (map f xs) = map f (strip_while (P \<circ> f) xs)"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    61
  by (simp add: strip_while_def rev_map dropWhile_map)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
    62
65388
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    63
lemma strip_while_dropWhile_commute:
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    64
  "strip_while P (dropWhile Q xs) = dropWhile Q (strip_while P xs)"
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    65
proof (induct xs)
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    66
  case Nil
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    67
  then show ?case
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    68
    by simp
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    69
next
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    70
  case (Cons x xs)
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    71
  show ?case
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    72
  proof (cases "\<forall>y\<in>set xs. P y")
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    73
    case True
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    74
    with dropWhile_append2 [of "rev xs"] show ?thesis
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    75
      by (auto simp add: strip_while_def dest: set_dropWhileD)
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    76
  next
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    77
    case False
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    78
    then obtain y where "y \<in> set xs" and "\<not> P y"
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    79
      by blast
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    80
    with Cons dropWhile_append3 [of P y "rev xs"] show ?thesis
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    81
      by (simp add: strip_while_def)
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    82
  qed
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    83
qed
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    84
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    85
lemma dropWhile_strip_while_commute:
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    86
  "dropWhile P (strip_while Q xs) = strip_while Q (dropWhile P xs)"
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    87
  by (simp add: strip_while_dropWhile_commute)
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
    88
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    89
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    90
definition no_leading :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    91
where
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    92
  "no_leading P xs \<longleftrightarrow> (xs \<noteq> [] \<longrightarrow> \<not> P (hd xs))"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    93
75008
43b3d5318d72 fixed dodgy intro! attributes
paulson <lp15@cam.ac.uk>
parents: 71420
diff changeset
    94
lemma no_leading_Nil [iff]:
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    95
  "no_leading P []"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    96
  by (simp add: no_leading_def)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    97
75008
43b3d5318d72 fixed dodgy intro! attributes
paulson <lp15@cam.ac.uk>
parents: 71420
diff changeset
    98
lemma no_leading_Cons [iff]:
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
    99
  "no_leading P (x # xs) \<longleftrightarrow> \<not> P x"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   100
  by (simp add: no_leading_def)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   101
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   102
lemma no_leading_append [simp]:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   103
  "no_leading P (xs @ ys) \<longleftrightarrow> no_leading P xs \<and> (xs = [] \<longrightarrow> no_leading P ys)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   104
  by (induct xs) simp_all
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   105
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   106
lemma no_leading_dropWhile [simp]:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   107
  "no_leading P (dropWhile P xs)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   108
  by (induct xs) simp_all
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   109
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   110
lemma dropWhile_eq_obtain_leading:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   111
  assumes "dropWhile P xs = ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   112
  obtains zs where "xs = zs @ ys" and "\<And>z. z \<in> set zs \<Longrightarrow> P z" and "no_leading P ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   113
proof -
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   114
  from assms have "\<exists>zs. xs = zs @ ys \<and> (\<forall>z \<in> set zs. P z) \<and> no_leading P ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   115
  proof (induct xs arbitrary: ys)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   116
    case Nil then show ?case by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   117
  next
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   118
    case (Cons x xs ys)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   119
    show ?case proof (cases "P x")
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   120
      case True with Cons.hyps [of ys] Cons.prems
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   121
      have "\<exists>zs. xs = zs @ ys \<and> (\<forall>a\<in>set zs. P a) \<and> no_leading P ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   122
        by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   123
      then obtain zs where "xs = zs @ ys" and "\<And>z. z \<in> set zs \<Longrightarrow> P z"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   124
        and *: "no_leading P ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   125
        by blast
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   126
      with True have "x # xs = (x # zs) @ ys" and "\<And>z. z \<in> set (x # zs) \<Longrightarrow> P z"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   127
        by auto
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   128
      with * show ?thesis
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   129
        by blast    next
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   130
      case False
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   131
      with Cons show ?thesis by (cases ys) simp_all
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   132
    qed
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   133
  qed
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   134
  with that show thesis
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   135
    by blast
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   136
qed
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   137
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   138
lemma dropWhile_idem_iff:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   139
  "dropWhile P xs = xs \<longleftrightarrow> no_leading P xs"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   140
  by (cases xs) (auto elim: dropWhile_eq_obtain_leading)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   141
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   142
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   143
abbreviation no_trailing :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   144
where
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   145
  "no_trailing P xs \<equiv> no_leading P (rev xs)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   146
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   147
lemma no_trailing_unfold:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   148
  "no_trailing P xs \<longleftrightarrow> (xs \<noteq> [] \<longrightarrow> \<not> P (last xs))"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   149
  by (induct xs) simp_all
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   150
75008
43b3d5318d72 fixed dodgy intro! attributes
paulson <lp15@cam.ac.uk>
parents: 71420
diff changeset
   151
lemma no_trailing_Nil [iff]:
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   152
  "no_trailing P []"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   153
  by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   154
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   155
lemma no_trailing_Cons [simp]:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   156
  "no_trailing P (x # xs) \<longleftrightarrow> no_trailing P xs \<and> (xs = [] \<longrightarrow> \<not> P x)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   157
  by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   158
65388
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   159
lemma no_trailing_append:
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   160
  "no_trailing P (xs @ ys) \<longleftrightarrow> no_trailing P ys \<and> (ys = [] \<longrightarrow> no_trailing P xs)"
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   161
  by (induct xs) simp_all
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   162
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   163
lemma no_trailing_append_Cons [simp]:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   164
  "no_trailing P (xs @ y # ys) \<longleftrightarrow> no_trailing P (y # ys)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   165
  by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   166
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   167
lemma no_trailing_strip_while [simp]:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   168
  "no_trailing P (strip_while P xs)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   169
  by (induct xs rule: rev_induct) simp_all
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   170
65388
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   171
lemma strip_while_idem [simp]:
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   172
  "no_trailing P xs \<Longrightarrow> strip_while P xs = xs"
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   173
  by (cases xs rule: rev_cases) simp_all
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   174
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   175
lemma strip_while_eq_obtain_trailing:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   176
  assumes "strip_while P xs = ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   177
  obtains zs where "xs = ys @ zs" and "\<And>z. z \<in> set zs \<Longrightarrow> P z" and "no_trailing P ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   178
proof -
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   179
  from assms have "rev (rev (dropWhile P (rev xs))) = rev ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   180
    by (simp add: strip_while_def)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   181
  then have "dropWhile P (rev xs) = rev ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   182
    by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   183
  then obtain zs where A: "rev xs = zs @ rev ys" and B: "\<And>z. z \<in> set zs \<Longrightarrow> P z"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   184
    and C: "no_trailing P ys"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   185
    using dropWhile_eq_obtain_leading by blast
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   186
  from A have "rev (rev xs) = rev (zs @ rev ys)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   187
    by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   188
  then have "xs = ys @ rev zs"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   189
    by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   190
  moreover from B have "\<And>z. z \<in> set (rev zs) \<Longrightarrow> P z"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   191
    by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   192
  ultimately show thesis using that C by blast
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   193
qed
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   194
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   195
lemma strip_while_idem_iff:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   196
  "strip_while P xs = xs \<longleftrightarrow> no_trailing P xs"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   197
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 60500
diff changeset
   198
  define ys where "ys = rev xs"
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   199
  moreover have "strip_while P (rev ys) = rev ys \<longleftrightarrow> no_trailing P (rev ys)"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   200
    by (simp add: dropWhile_idem_iff)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   201
  ultimately show ?thesis by simp
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   202
qed
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   203
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   204
lemma no_trailing_map:
65388
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   205
  "no_trailing P (map f xs) \<longleftrightarrow> no_trailing (P \<circ> f) xs"
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   206
  by (simp add: last_map no_trailing_unfold)
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   207
65388
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   208
lemma no_trailing_drop [simp]:
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   209
  "no_trailing P (drop n xs)" if "no_trailing P xs"
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   210
proof -
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   211
  from that have "no_trailing P (take n xs @ drop n xs)"
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   212
    by simp
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   213
  then show ?thesis
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   214
    by (simp only: no_trailing_append)
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   215
qed
a8d868477bc0 more on lists
haftmann
parents: 63540
diff changeset
   216
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   217
lemma no_trailing_upt [simp]:
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   218
  "no_trailing P [n..<m] \<longleftrightarrow> (n < m \<longrightarrow> \<not> P (m - 1))"
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   219
  by (auto simp add: no_trailing_unfold)
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   220
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   221
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   222
definition nth_default :: "'a \<Rightarrow> 'a list \<Rightarrow> nat \<Rightarrow> 'a"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   223
where
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   224
  "nth_default dflt xs n = (if n < length xs then xs ! n else dflt)"
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   225
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   226
lemma nth_default_nth:
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   227
  "n < length xs \<Longrightarrow> nth_default dflt xs n = xs ! n"
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   228
  by (simp add: nth_default_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   229
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   230
lemma nth_default_beyond:
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   231
  "length xs \<le> n \<Longrightarrow> nth_default dflt xs n = dflt"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   232
  by (simp add: nth_default_def)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   233
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   234
lemma nth_default_Nil [simp]:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   235
  "nth_default dflt [] n = dflt"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   236
  by (simp add: nth_default_def)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   237
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   238
lemma nth_default_Cons:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   239
  "nth_default dflt (x # xs) n = (case n of 0 \<Rightarrow> x | Suc n' \<Rightarrow> nth_default dflt xs n')"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   240
  by (simp add: nth_default_def split: nat.split)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   241
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   242
lemma nth_default_Cons_0 [simp]:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   243
  "nth_default dflt (x # xs) 0 = x"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   244
  by (simp add: nth_default_Cons)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   245
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   246
lemma nth_default_Cons_Suc [simp]:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   247
  "nth_default dflt (x # xs) (Suc n) = nth_default dflt xs n"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   248
  by (simp add: nth_default_Cons)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   249
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   250
lemma nth_default_replicate_dflt [simp]:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   251
  "nth_default dflt (replicate n dflt) m = dflt"
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   252
  by (simp add: nth_default_def)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   253
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   254
lemma nth_default_append:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   255
  "nth_default dflt (xs @ ys) n =
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   256
    (if n < length xs then nth xs n else nth_default dflt ys (n - length xs))"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   257
  by (auto simp add: nth_default_def nth_append)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   258
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   259
lemma nth_default_append_trailing [simp]:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   260
  "nth_default dflt (xs @ replicate n dflt) = nth_default dflt xs"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   261
  by (simp add: fun_eq_iff nth_default_append) (simp add: nth_default_def)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   262
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   263
lemma nth_default_snoc_default [simp]:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   264
  "nth_default dflt (xs @ [dflt]) = nth_default dflt xs"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   265
  by (auto simp add: nth_default_def fun_eq_iff nth_append)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   266
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   267
lemma nth_default_eq_dflt_iff:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   268
  "nth_default dflt xs k = dflt \<longleftrightarrow> (k < length xs \<longrightarrow> xs ! k = dflt)"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   269
  by (simp add: nth_default_def)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   270
67730
f91c437f6f68 new lemma
haftmann
parents: 67399
diff changeset
   271
lemma nth_default_take_eq:
f91c437f6f68 new lemma
haftmann
parents: 67399
diff changeset
   272
  "nth_default dflt (take m xs) n =
f91c437f6f68 new lemma
haftmann
parents: 67399
diff changeset
   273
    (if n < m then nth_default dflt xs n else dflt)"
f91c437f6f68 new lemma
haftmann
parents: 67399
diff changeset
   274
  by (simp add: nth_default_def)
f91c437f6f68 new lemma
haftmann
parents: 67399
diff changeset
   275
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   276
lemma in_enumerate_iff_nth_default_eq:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   277
  "x \<noteq> dflt \<Longrightarrow> (n, x) \<in> set (enumerate 0 xs) \<longleftrightarrow> nth_default dflt xs n = x"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   278
  by (auto simp add: nth_default_def in_set_conv_nth enumerate_eq_zip)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   279
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   280
lemma last_conv_nth_default:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   281
  assumes "xs \<noteq> []"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   282
  shows "last xs = nth_default dflt xs (length xs - 1)"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   283
  using assms by (simp add: nth_default_def last_conv_nth)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   284
  
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   285
lemma nth_default_map_eq:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   286
  "f dflt' = dflt \<Longrightarrow> nth_default dflt (map f xs) n = f (nth_default dflt' xs n)"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   287
  by (simp add: nth_default_def)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   288
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   289
lemma finite_nth_default_neq_default [simp]:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   290
  "finite {k. nth_default dflt xs k \<noteq> dflt}"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   291
  by (simp add: nth_default_def)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   292
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   293
lemma sorted_list_of_set_nth_default:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   294
  "sorted_list_of_set {k. nth_default dflt xs k \<noteq> dflt} = map fst (filter (\<lambda>(_, x). x \<noteq> dflt) (enumerate 0 xs))"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   295
  by (rule sorted_distinct_set_unique) (auto simp add: nth_default_def in_set_conv_nth
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   296
    sorted_filter distinct_map_filter enumerate_eq_zip intro: rev_image_eqI)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   297
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   298
lemma map_nth_default:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   299
  "map (nth_default x xs) [0..<length xs] = xs"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   300
proof -
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   301
  have *: "map (nth_default x xs) [0..<length xs] = map (List.nth xs) [0..<length xs]"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   302
    by (rule map_cong) (simp_all add: nth_default_nth)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   303
  show ?thesis by (simp add: * map_nth)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   304
qed
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   305
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   306
lemma range_nth_default [simp]:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   307
  "range (nth_default dflt xs) = insert dflt (set xs)"
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   308
  by (auto simp add: nth_default_def [abs_def] in_set_conv_nth)
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   309
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   310
lemma nth_strip_while:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   311
  assumes "n < length (strip_while P xs)"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   312
  shows "strip_while P xs ! n = xs ! n"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   313
proof -
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   314
  have "length (dropWhile P (rev xs)) + length (takeWhile P (rev xs)) = length xs"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   315
    by (subst add.commute)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   316
      (simp add: arg_cong [where f=length, OF takeWhile_dropWhile_id, unfolded length_append])
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   317
  then show ?thesis using assms
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   318
    by (simp add: strip_while_def rev_nth dropWhile_nth)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   319
qed
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   320
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   321
lemma length_strip_while_le:
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   322
  "length (strip_while P xs) \<le> length xs"
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   323
  unfolding strip_while_def o_def length_rev
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   324
  by (subst (2) length_rev[symmetric])
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   325
    (simp add: strip_while_def length_dropWhile_le del: length_rev)
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   326
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   327
lemma nth_default_strip_while_dflt [simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 65388
diff changeset
   328
  "nth_default dflt (strip_while ((=) dflt) xs) = nth_default dflt xs"
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   329
  by (induct xs rule: rev_induct) auto
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   330
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   331
lemma nth_default_eq_iff:
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   332
  "nth_default dflt xs = nth_default dflt ys
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   333
     \<longleftrightarrow> strip_while (HOL.eq dflt) xs = strip_while (HOL.eq dflt) ys" (is "?P \<longleftrightarrow> ?Q")
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   334
proof
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   335
  let ?xs = "strip_while (HOL.eq dflt) xs" and ?ys = "strip_while (HOL.eq dflt) ys"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   336
  assume ?P
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   337
  then have eq: "nth_default dflt ?xs = nth_default dflt ?ys"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   338
    by simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   339
  have len: "length ?xs = length ?ys"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   340
  proof (rule ccontr)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   341
    assume len: "length ?xs \<noteq> length ?ys"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   342
    { fix xs ys :: "'a list"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   343
      let ?xs = "strip_while (HOL.eq dflt) xs" and ?ys = "strip_while (HOL.eq dflt) ys"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   344
      assume eq: "nth_default dflt ?xs = nth_default dflt ?ys"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   345
      assume len: "length ?xs < length ?ys"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   346
      then have "length ?ys > 0" by arith
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   347
      then have "?ys \<noteq> []" by simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   348
      with last_conv_nth_default [of ?ys dflt]
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   349
      have "last ?ys = nth_default dflt ?ys (length ?ys - 1)"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   350
        by auto
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   351
      moreover from \<open>?ys \<noteq> []\<close> no_trailing_strip_while [of "HOL.eq dflt" ys]
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   352
        have "last ?ys \<noteq> dflt" by (simp add: no_trailing_unfold)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   353
      ultimately have "nth_default dflt ?xs (length ?ys - 1) \<noteq> dflt"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   354
        using eq by simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   355
      moreover from len have "length ?ys - 1 \<ge> length ?xs" by simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   356
      ultimately have False by (simp only: nth_default_beyond) simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   357
    } 
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   358
    from this [of xs ys] this [of ys xs] len eq show False
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   359
      by (auto simp only: linorder_class.neq_iff)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   360
  qed
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   361
  then show ?Q
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   362
  proof (rule nth_equalityI [rule_format])
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   363
    fix n
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63040
diff changeset
   364
    assume n: "n < length ?xs"
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63040
diff changeset
   365
    with len have "n < length ?ys"
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   366
      by simp
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63040
diff changeset
   367
    with n have xs: "nth_default dflt ?xs n = ?xs ! n"
58437
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   368
      and ys: "nth_default dflt ?ys n = ?ys ! n"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   369
      by (simp_all only: nth_default_nth)
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   370
    with eq show "?xs ! n = ?ys ! n"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   371
      by simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   372
  qed
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   373
next
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   374
  assume ?Q
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   375
  then have "nth_default dflt (strip_while (HOL.eq dflt) xs) = nth_default dflt (strip_while (HOL.eq dflt) ys)"
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   376
    by simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   377
  then show ?P
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   378
    by simp
8d124c73c37a added lemmas
haftmann
parents: 58295
diff changeset
   379
qed
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   380
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   381
lemma nth_default_map2:
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   382
  \<open>nth_default d (map2 f xs ys) n = f (nth_default d1 xs n) (nth_default d2 ys n)\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   383
    if \<open>length xs = length ys\<close> and \<open>f d1 d2 = d\<close> for bs cs
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   384
using that proof (induction xs ys arbitrary: n rule: list_induct2)
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   385
  case Nil
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   386
  then show ?case
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   387
    by simp
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   388
next
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   389
  case (Cons x xs y ys)
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   390
  then show ?case
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   391
    by (cases n) simp_all
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   392
qed
572ab9e64e18 simplified logical constructions
haftmann
parents: 67730
diff changeset
   393
58199
5fbe474b5da8 explicit theory with additional, less commonly used list operations
haftmann
parents:
diff changeset
   394
end
58295
c8a8e7c37986 more material on lists
haftmann
parents: 58199
diff changeset
   395