prio/Moment.thy
author urbanc
Tue, 06 Mar 2012 11:30:45 +0000
changeset 337 f9d54f49c808
parent 336 f9e0d3274c14
child 339 b3add51e2e0f
permissions -rw-r--r--
typo
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     1
theory Moment
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     2
imports Main
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     3
begin
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     4
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     5
fun firstn :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     6
where
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     7
  "firstn 0 s = []" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     8
  "firstn (Suc n) [] = []" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     9
  "firstn (Suc n) (e#s) = e#(firstn n s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    10
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    11
fun restn :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    12
where "restn n s = rev (firstn (length s - n) (rev s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    13
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    14
definition moment :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    15
where "moment n s = rev (firstn n (rev s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    16
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    17
definition restm :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    18
where "restm n s = rev (restn n (rev s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    19
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    20
definition from_to :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    21
  where "from_to i j s = firstn (j - i) (restn i s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    22
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    23
definition down_to :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    24
where "down_to j i s = rev (from_to i j (rev s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    25
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    26
(*
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    27
value "down_to 6 2 [10, 9, 8, 7, 6, 5, 4, 3, 2, 1, 0]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    28
value "from_to 2 6 [0, 1, 2, 3, 4, 5, 6, 7]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    29
*)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    30
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    31
lemma length_eq_elim_l: "\<lbrakk>length xs = length ys; xs@us = ys@vs\<rbrakk> \<Longrightarrow> xs = ys \<and> us = vs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    32
  by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    33
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    34
lemma length_eq_elim_r: "\<lbrakk>length us = length vs; xs@us = ys@vs\<rbrakk> \<Longrightarrow> xs = ys \<and> us = vs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    35
  by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    36
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    37
lemma firstn_nil [simp]: "firstn n [] = []"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    38
  by (cases n, simp+)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    39
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    40
(*
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    41
value "from_to 0 2 [0, 1, 2, 3, 4, 5, 6, 7, 8, 9] @ 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    42
       from_to 2 5 [0, 1, 2, 3, 4, 5, 6, 7, 8, 9]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    43
*)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    44
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    45
lemma firstn_le: "\<And> n s'. n \<le> length s \<Longrightarrow> firstn n (s@s') = firstn n s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    46
proof (induct s, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    47
  fix a s n s'
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    48
  assume ih: "\<And>n s'. n \<le> length s \<Longrightarrow> firstn n (s @ s') = firstn n s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    49
  and le_n: " n \<le> length (a # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    50
  show "firstn n ((a # s) @ s') = firstn n (a # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    51
  proof(cases n, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    52
    fix k
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    53
    assume eq_n: "n = Suc k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    54
    with le_n have "k \<le> length s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    55
    from ih [OF this] and eq_n
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    56
    show "firstn n ((a # s) @ s') = firstn n (a # s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    57
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    58
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    59
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    60
lemma firstn_ge [simp]: "\<And>n. length s \<le> n \<Longrightarrow> firstn n s = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    61
proof(induct s, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    62
  fix a s n
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    63
  assume ih: "\<And>n. length s \<le> n \<Longrightarrow> firstn n s = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    64
    and le: "length (a # s) \<le> n"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    65
  show "firstn n (a # s) = a # s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    66
  proof(cases n)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    67
    assume eq_n: "n = 0" with le show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    68
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    69
    fix k
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    70
    assume eq_n: "n = Suc k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    71
    with le have le_k: "length s \<le> k" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    72
    from ih [OF this] have "firstn k s = s" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    73
    from eq_n and this
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    74
    show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    75
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    76
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    77
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    78
lemma firstn_eq [simp]: "firstn (length s) s = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    79
  by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    80
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    81
lemma firstn_restn_s: "(firstn n (s::'a list)) @ (restn n s) = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    82
proof(induct n arbitrary:s, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    83
  fix n s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    84
  assume ih: "\<And>t. firstn n (t::'a list) @ restn n t = t"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    85
  show "firstn (Suc n) (s::'a list) @ restn (Suc n) s = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    86
  proof(cases s, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    87
    fix x xs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    88
    assume eq_s: "s = x#xs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    89
    show "firstn (Suc n) s @ restn (Suc n) s = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    90
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    91
      have "firstn (Suc n) s @ restn (Suc n) s =  x # (firstn n xs @ restn n xs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    92
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    93
        from eq_s have "firstn (Suc n) s =  x # firstn n xs" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    94
        moreover have "restn (Suc n) s = restn n xs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    95
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    96
          from eq_s have "restn (Suc n) s = rev (firstn (length xs - n) (rev xs @ [x]))" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    97
          also have "\<dots> = restn n xs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    98
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    99
            have "(firstn (length xs - n) (rev xs @ [x])) = (firstn (length xs - n) (rev xs))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   100
              by(rule firstn_le, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   101
            hence "rev (firstn (length xs - n) (rev xs @ [x])) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   102
              rev (firstn (length xs - n) (rev xs))" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   103
            also have "\<dots> = rev (firstn (length (rev xs) - n) (rev xs))" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   104
            finally show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   105
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   106
          finally show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   107
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   108
        ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   109
      qed with ih eq_s show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   110
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   111
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   112
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   113
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   114
lemma moment_restm_s: "(restm n s)@(moment n s) = s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   115
by (metis firstn_restn_s moment_def restm_def rev_append rev_rev_ident)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   116
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   117
declare restn.simps [simp del] firstn.simps[simp del]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   118
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   119
lemma length_firstn_ge: "length s \<le> n \<Longrightarrow> length (firstn n s) = length s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   120
by (metis firstn_ge)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   121
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   122
lemma length_firstn_le: "n \<le> length s \<Longrightarrow> length (firstn n s) = n"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   123
proof(induct n arbitrary:s, simp add:firstn.simps)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   124
  case (Suc k)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   125
  assume ih: "\<And>s. k \<le> length (s::'a list) \<Longrightarrow> length (firstn k s) = k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   126
    and le: "Suc k \<le> length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   127
  show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   128
  proof(cases s)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   129
    case Nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   130
    from Nil and le show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   131
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   132
    case (Cons x xs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   133
    from le and Cons have "k \<le> length xs" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   134
    from ih [OF this] have "length (firstn k xs) = k" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   135
    moreover from Cons have "length (firstn (Suc k) s) = Suc (length (firstn k xs))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   136
      by (simp add:firstn.simps)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   137
    ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   138
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   139
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   140
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   141
lemma app_firstn_restn: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   142
  fixes s1 s2
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   143
  shows "s1 = firstn (length s1) (s1 @ s2) \<and> s2 = restn (length s1) (s1 @ s2)"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   144
by (metis append_eq_conv_conj firstn_ge firstn_le firstn_restn_s le_refl)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   145
lemma length_moment_le:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   146
  fixes k s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   147
  assumes le_k: "k \<le> length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   148
  shows "length (moment k s) = k"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   149
by (metis assms length_firstn_le length_rev moment_def)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   150
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   151
lemma app_moment_restm: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   152
  fixes s1 s2
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   153
  shows "s1 = restm (length s2) (s1 @ s2) \<and> s2 = moment (length s2) (s1 @ s2)"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   154
by (metis app_firstn_restn length_rev moment_def restm_def rev_append rev_rev_ident)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   155
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   156
lemma length_moment_ge:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   157
  fixes k s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   158
  assumes le_k: "length s \<le> k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   159
  shows "length (moment k s) = (length s)"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   160
by (metis assms firstn_ge length_rev moment_def)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   161
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   162
lemma length_firstn: "(length (firstn n s) = length s) \<or> (length (firstn n s) = n)"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   163
by (metis length_firstn_ge length_firstn_le nat_le_linear)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   164
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   165
lemma firstn_conc: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   166
  fixes m n
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   167
  assumes le_mn: "m \<le> n"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   168
  shows "firstn m s = firstn m (firstn n  s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   169
proof(cases "m \<le> length s")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   170
  case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   171
  have "s = (firstn n s) @ (restn n s)" by (simp add:firstn_restn_s)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   172
  hence "firstn m s = firstn m \<dots>" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   173
  also have "\<dots> = firstn m (firstn n s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   174
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   175
    from length_firstn [of n s]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   176
    have "m \<le> length (firstn n s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   177
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   178
      assume "length (firstn n s) = length s" with True show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   179
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   180
      assume "length (firstn n s) = n " with le_mn show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   181
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   182
    from firstn_le [OF this, of "restn n s"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   183
    show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   184
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   185
  finally show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   186
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   187
  case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   188
  from False and le_mn have "length s \<le> n"  by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   189
  from firstn_ge [OF this] show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   190
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   191
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   192
lemma restn_conc: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   193
  fixes i j k s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   194
  assumes eq_k: "j + i = k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   195
  shows "restn k s = restn j (restn i s)"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   196
by (metis app_moment_restm append_take_drop_id assms drop_drop length_drop moment_def restn.simps)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   197
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   198
(*
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   199
value "down_to 2 0 [5, 4, 3, 2, 1, 0]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   200
value "moment 2 [5, 4, 3, 2, 1, 0]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   201
*)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   202
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   203
lemma from_to_firstn: "from_to 0 k s = firstn k s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   204
by (simp add:from_to_def restn.simps)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   205
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   206
lemma moment_app [simp]:
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   207
  assumes ile: "i \<le> length s"
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   208
  shows "moment i (s'@s) = moment i s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   209
by (metis assms firstn_le length_rev moment_def rev_append)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   210
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   211
lemma moment_eq [simp]: "moment (length s) (s'@s) = s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   212
by (metis app_moment_restm)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   213
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   214
lemma moment_ge [simp]: "length s \<le> n \<Longrightarrow> moment n s = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   215
  by (unfold moment_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   216
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   217
lemma moment_zero [simp]: "moment 0 s = []"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   218
  by (simp add:moment_def firstn.simps)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   219
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   220
lemma p_split_gen: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   221
  "\<lbrakk>Q s; \<not> Q (moment k s)\<rbrakk> \<Longrightarrow>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   222
  (\<exists> i. i < length s \<and> k \<le> i \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   223
proof (induct s, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   224
  fix a s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   225
  assume ih: "\<lbrakk>Q s; \<not> Q (moment k s)\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   226
           \<Longrightarrow> \<exists>i<length s. k \<le> i \<and> \<not> Q (moment i s) \<and> (\<forall>i'>i. Q (moment i' s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   227
    and nq: "\<not> Q (moment k (a # s))" and qa: "Q (a # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   228
  have le_k: "k \<le> length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   229
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   230
    { assume "length s < k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   231
      hence "length (a#s) \<le> k" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   232
      from moment_ge [OF this] and nq and qa
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   233
      have "False" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   234
    } thus ?thesis by arith
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   235
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   236
  have nq_k: "\<not> Q (moment k s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   237
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   238
    have "moment k (a#s) = moment k s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   239
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   240
      from moment_app [OF le_k, of "[a]"] show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   241
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   242
    with nq show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   243
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   244
  show "\<exists>i<length (a # s). k \<le> i \<and> \<not> Q (moment i (a # s)) \<and> (\<forall>i'>i. Q (moment i' (a # s)))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   245
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   246
    { assume "Q s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   247
      from ih [OF this nq_k]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   248
      obtain i where lti: "i < length s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   249
        and nq: "\<not> Q (moment i s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   250
        and rst: "\<forall>i'>i. Q (moment i' s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   251
        and lki: "k \<le> i" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   252
      have ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   253
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   254
        from lti have "i < length (a # s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   255
        moreover have " \<not> Q (moment i (a # s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   256
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   257
          from lti have "i \<le> (length s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   258
          from moment_app [OF this, of "[a]"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   259
          have "moment i (a # s) = moment i s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   260
          with nq show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   261
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   262
        moreover have " (\<forall>i'>i. Q (moment i' (a # s)))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   263
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   264
          {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   265
            fix i'
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   266
            assume lti': "i < i'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   267
            have "Q (moment i' (a # s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   268
            proof(cases "length (a#s) \<le> i'")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   269
              case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   270
              from True have "moment i' (a#s) = a#s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   271
              with qa show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   272
            next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   273
              case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   274
              from False have "i' \<le> length s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   275
              from moment_app [OF this, of "[a]"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   276
              have "moment i' (a#s) = moment i' s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   277
              with rst lti' show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   278
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   279
          } thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   280
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   281
        moreover note lki
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   282
        ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   283
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   284
    } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   285
      assume ns: "\<not> Q s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   286
      have ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   287
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   288
        let ?i = "length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   289
        have "\<not> Q (moment ?i (a#s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   290
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   291
          have "?i \<le> length s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   292
          from moment_app [OF this, of "[a]"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   293
          have "moment ?i (a#s) = moment ?i s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   294
          moreover have "\<dots> = s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   295
          ultimately show ?thesis using ns by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   296
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   297
        moreover have "\<forall> i' > ?i. Q (moment i' (a#s))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   298
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   299
          { fix i'
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   300
            assume "i' > ?i"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   301
            hence "length (a#s) \<le> i'" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   302
            from moment_ge [OF this] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   303
            have " moment i' (a # s) = a # s" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   304
            with qa have "Q (moment i' (a#s))" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   305
          } thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   306
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   307
        moreover have "?i < length (a#s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   308
        moreover note le_k
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   309
        ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   310
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   311
    } ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   312
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   313
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   314
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   315
lemma p_split: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   316
  "\<And> s Q. \<lbrakk>Q s; \<not> Q []\<rbrakk> \<Longrightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   317
       (\<exists> i. i < length s \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   318
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   319
  fix s Q
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   320
  assume qs: "Q s" and nq: "\<not> Q []"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   321
  from nq have "\<not> Q (moment 0 s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   322
  from p_split_gen [of Q s 0, OF qs this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   323
  show "(\<exists> i. i < length s \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   324
    by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   325
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   326
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   327
lemma moment_plus: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   328
  "Suc i \<le> length s \<Longrightarrow> moment (Suc i) s = (hd (moment (Suc i) s)) # (moment i s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   329
proof(induct s, simp+)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   330
  fix a s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   331
  assume ih: "Suc i \<le> length s \<Longrightarrow> moment (Suc i) s = hd (moment (Suc i) s) # moment i s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   332
    and le_i: "i \<le> length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   333
  show "moment (Suc i) (a # s) = hd (moment (Suc i) (a # s)) # moment i (a # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   334
  proof(cases "i= length s")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   335
    case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   336
    hence "Suc i = length (a#s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   337
    with moment_eq have "moment (Suc i) (a#s) = a#s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   338
    moreover have "moment i (a#s) = s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   339
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   340
      from moment_app [OF le_i, of "[a]"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   341
      and True show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   342
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   343
    ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   344
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   345
    case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   346
    from False and le_i have lti: "i < length s" by arith
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   347
    hence les_i: "Suc i \<le> length s" by arith
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   348
    show ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   349
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   350
      from moment_app [OF les_i, of "[a]"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   351
      have "moment (Suc i) (a # s) = moment (Suc i) s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   352
      moreover have "moment i (a#s) = moment i s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   353
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   354
        from lti have "i \<le> length s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   355
        from moment_app [OF this, of "[a]"] show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   356
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   357
      moreover note ih [OF les_i]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   358
      ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   359
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   360
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   361
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   362
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   363
lemma from_to_conc:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   364
  fixes i j k s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   365
  assumes le_ij: "i \<le> j"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   366
  and le_jk: "j \<le> k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   367
  shows "from_to i j s @ from_to j k s = from_to i k s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   368
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   369
  let ?ris = "restn i s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   370
  have "firstn (j - i) (restn i s) @ firstn (k - j) (restn j s) =
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   371
         firstn (k - i) (restn i s)" (is "?x @ ?y = ?z")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   372
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   373
    let "firstn (k-j) ?u" = "?y"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   374
    let ?rst = " restn (k - j) (restn (j - i) ?ris)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   375
    let ?rst' = "restn (k - i) ?ris"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   376
    have "?u = restn (j-i) ?ris"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   377
    proof(rule restn_conc)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   378
      from le_ij show "j - i + i = j" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   379
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   380
    hence "?x @ ?y = ?x @ firstn (k-j) (restn (j-i) ?ris)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   381
    moreover have "firstn (k - j) (restn (j - i) (restn i s)) @ ?rst = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   382
                      restn (j-i) ?ris" by (simp add:firstn_restn_s)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   383
    ultimately have "?x @ ?y @ ?rst = ?x @ (restn (j-i) ?ris)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   384
    also have "\<dots> = ?ris" by (simp add:firstn_restn_s)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   385
    finally have "?x @ ?y @ ?rst = ?ris" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   386
    moreover have "?z @ ?rst = ?ris"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   387
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   388
      have "?z @ ?rst' = ?ris" by (simp add:firstn_restn_s)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   389
      moreover have "?rst' = ?rst"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   390
      proof(rule restn_conc)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   391
        from le_ij le_jk show "k - j + (j - i) = k - i" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   392
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   393
      ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   394
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   395
    ultimately have "?x @ ?y @ ?rst = ?z @ ?rst" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   396
    thus ?thesis by auto    
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   397
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   398
  thus ?thesis by (simp only:from_to_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   399
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   400
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   401
lemma down_to_conc:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   402
  fixes i j k s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   403
  assumes le_ij: "i \<le> j"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   404
  and le_jk: "j \<le> k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   405
  shows "down_to k j s @ down_to j i s = down_to k i s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   406
by (metis down_to_def from_to_conc le_ij le_jk rev_append)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   407
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   408
lemma restn_ge:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   409
  fixes s k
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   410
  assumes le_k: "length s \<le> k"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   411
  shows "restn k s = []"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   412
by (metis assms diff_is_0_eq moment_def moment_zero restn.simps)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   413
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   414
lemma from_to_ge: "length s \<le> k \<Longrightarrow> from_to k j s = []"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   415
by (metis firstn_nil from_to_def restn_ge)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   416
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   417
(*
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   418
value "from_to 2 5 [0, 1, 2, 3, 4]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   419
value "restn 2  [0, 1, 2, 3, 4]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   420
*)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   421
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   422
lemma from_to_restn: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   423
  fixes k j s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   424
  assumes le_j: "length s \<le> j"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   425
  shows "from_to k j s = restn k s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   426
by (metis app_moment_restm append_Nil2 append_take_drop_id assms diff_is_0_eq' drop_take firstn_restn_s from_to_def length_drop moment_def moment_zero restn.simps)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   427
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   428
lemma down_to_moment: "down_to k 0 s = moment k s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   429
by (metis down_to_def from_to_firstn moment_def)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   430
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   431
lemma down_to_restm:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   432
  assumes le_s: "length s \<le> j"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   433
  shows "down_to j k s = restm k s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   434
by (metis assms down_to_def from_to_restn length_rev restm_def)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   435
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   436
lemma moment_split: "moment (m+i) s = down_to (m+i) i s @down_to i 0 s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   437
by (metis down_to_conc down_to_moment le0 le_add1 nat_add_commute)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   438
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   439
lemma length_restn: "length (restn i s) = length s - i"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   440
by (metis diff_le_self length_firstn_le length_rev restn.simps)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   441
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   442
lemma length_from_to_in:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   443
  fixes i j s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   444
  assumes le_ij: "i \<le> j"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   445
  and le_j: "j \<le> length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   446
  shows "length (from_to i j s) = j - i"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   447
by (metis diff_le_mono from_to_def le_j length_firstn_le length_restn)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   448
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   449
lemma firstn_restn_from_to: "from_to i (m + i) s = firstn m (restn i s)"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   450
by (metis diff_add_inverse2 from_to_def)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   451
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   452
lemma down_to_moment_restm:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   453
  fixes m i s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   454
  shows "down_to (m + i) i s = moment m (restm i s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   455
  by (simp add:firstn_restn_from_to down_to_def moment_def restm_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   456
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   457
lemma moment_plus_split:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   458
  fixes m i s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   459
  shows "moment (m + i) s = moment m (restm i s) @ moment i s"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   460
by (metis down_to_moment down_to_moment_restm moment_split)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   461
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   462
lemma length_restm: "length (restm i s) = length s - i"
336
f9e0d3274c14 fixed typo
urbanc
parents: 262
diff changeset
   463
by (metis length_restn length_rev restm_def)
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   464
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   465
end