Moment.thy~
author zhangx
Thu, 28 Jan 2016 16:36:46 +0800
changeset 89 2056d9f481e2
parent 81 c495eb16beb6
permissions -rw-r--r--
Slightly modified ExtGG.thy and PrioG.thy.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
81
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     1
theory Moment
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     2
imports Main
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     3
begin
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     4
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     5
definition moment :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     6
where "moment n s = rev (take n (rev s))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     7
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     8
value "moment 3 [0, 1, 2, 3, 4, 5, 6, 7, 8, 9::int]"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
     9
value "moment 2 [5, 4, 3, 2, 1, 0::int]"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    10
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    11
(*
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    12
lemma length_moment_le:
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    13
  assumes le_k: "k \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    14
  shows "length (moment k s) = k"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    15
using le_k unfolding moment_def by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    16
*)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    17
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    18
(*
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    19
lemma length_moment_ge:
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    20
  assumes le_k: "length s \<le> k"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    21
  shows "length (moment k s) = (length s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    22
using assms unfolding moment_def by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    23
*)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    24
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    25
lemma moment_app [simp]:
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    26
  assumes ile: "i \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    27
  shows "moment i (s' @ s) = moment i s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    28
using assms unfolding moment_def by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    29
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    30
lemma moment_eq [simp]: "moment (length s) (s' @ s) = s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    31
  unfolding moment_def by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    32
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    33
lemma moment_ge [simp]: "length s \<le> n \<Longrightarrow> moment n s = s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    34
  by (unfold moment_def, simp)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    35
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    36
lemma moment_zero [simp]: "moment 0 s = []"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    37
  by (simp add:moment_def)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    38
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    39
lemma p_split_gen: 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    40
  "\<lbrakk>Q s; \<not> Q (moment k s)\<rbrakk> \<Longrightarrow>
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    41
  (\<exists> i. i < length s \<and> k \<le> i \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    42
proof (induct s, simp)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    43
  fix a s
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    44
  assume ih: "\<lbrakk>Q s; \<not> Q (moment k s)\<rbrakk>
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    45
           \<Longrightarrow> \<exists>i<length s. k \<le> i \<and> \<not> Q (moment i s) \<and> (\<forall>i'>i. Q (moment i' s))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    46
    and nq: "\<not> Q (moment k (a # s))" and qa: "Q (a # s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    47
  have le_k: "k \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    48
  proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    49
    { assume "length s < k"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    50
      hence "length (a#s) \<le> k" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    51
      from moment_ge [OF this] and nq and qa
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    52
      have "False" by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    53
    } thus ?thesis by arith
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    54
  qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    55
  have nq_k: "\<not> Q (moment k s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    56
  proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    57
    have "moment k (a#s) = moment k s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    58
    proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    59
      from moment_app [OF le_k, of "[a]"] show ?thesis by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    60
    qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    61
    with nq show ?thesis by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    62
  qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    63
  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)))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    64
  proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    65
    { assume "Q s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    66
      from ih [OF this nq_k]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    67
      obtain i where lti: "i < length s" 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    68
        and nq: "\<not> Q (moment i s)" 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    69
        and rst: "\<forall>i'>i. Q (moment i' s)" 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    70
        and lki: "k \<le> i" by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    71
      have ?thesis 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    72
      proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    73
        from lti have "i < length (a # s)" by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    74
        moreover have " \<not> Q (moment i (a # s))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    75
        proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    76
          from lti have "i \<le> (length s)" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    77
          from moment_app [OF this, of "[a]"]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    78
          have "moment i (a # s) = moment i s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    79
          with nq show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    80
        qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    81
        moreover have " (\<forall>i'>i. Q (moment i' (a # s)))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    82
        proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    83
          {
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    84
            fix i'
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    85
            assume lti': "i < i'"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    86
            have "Q (moment i' (a # s))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    87
            proof(cases "length (a#s) \<le> i'")
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    88
              case True
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    89
              from True have "moment i' (a#s) = a#s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    90
              with qa show ?thesis by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    91
            next
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    92
              case False
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    93
              from False have "i' \<le> length s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    94
              from moment_app [OF this, of "[a]"]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    95
              have "moment i' (a#s) = moment i' s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    96
              with rst lti' show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    97
            qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    98
          } thus ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
    99
        qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   100
        moreover note lki
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   101
        ultimately show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   102
      qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   103
    } moreover {
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   104
      assume ns: "\<not> Q s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   105
      have ?thesis
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   106
      proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   107
        let ?i = "length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   108
        have "\<not> Q (moment ?i (a#s))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   109
        proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   110
          have "?i \<le> length s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   111
          from moment_app [OF this, of "[a]"]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   112
          have "moment ?i (a#s) = moment ?i s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   113
          moreover have "\<dots> = s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   114
          ultimately show ?thesis using ns by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   115
        qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   116
        moreover have "\<forall> i' > ?i. Q (moment i' (a#s))" 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   117
        proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   118
          { fix i'
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   119
            assume "i' > ?i"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   120
            hence "length (a#s) \<le> i'" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   121
            from moment_ge [OF this] 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   122
            have " moment i' (a # s) = a # s" .
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   123
            with qa have "Q (moment i' (a#s))" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   124
          } thus ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   125
        qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   126
        moreover have "?i < length (a#s)" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   127
        moreover note le_k
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   128
        ultimately show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   129
      qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   130
    } ultimately show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   131
  qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   132
qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   133
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   134
lemma p_split: 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   135
  "\<lbrakk>Q s; \<not> Q []\<rbrakk> \<Longrightarrow> 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   136
       (\<exists> i. i < length s \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   137
proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   138
  fix s Q
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   139
  assume qs: "Q s" and nq: "\<not> Q []"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   140
  from nq have "\<not> Q (moment 0 s)" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   141
  from p_split_gen [of Q s 0, OF qs this]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   142
  show "(\<exists> i. i < length s \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   143
    by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   144
qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   145
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   146
lemma moment_Suc_tl:
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   147
  assumes "Suc i \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   148
  shows "tl (moment (Suc i) s) = moment i s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   149
  using assms unfolding moment_def rev_take
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   150
by (simp, metis Suc_diff_le diff_Suc_Suc drop_Suc tl_drop)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   151
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   152
lemma moment_plus:
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   153
  assumes "Suc i \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   154
  shows "(moment (Suc i) s) = (hd (moment (Suc i) s)) # (moment i s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   155
proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   156
  have "(moment (Suc i) s) \<noteq> []"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   157
  using assms by (auto simp add: moment_def)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   158
  hence "(moment (Suc i) s) = (hd (moment (Suc i) s)) #  tl (moment (Suc i) s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   159
    by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   160
<<<<<<< local
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   161
  have "moment (i + length s) (t @ s) = moment i t @ moment (length s) (t @ s)" 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   162
    by (simp add: moment_def)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   163
  with moment_app show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   164
qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   165
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   166
lemma moment_Suc_tl:
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   167
  assumes "Suc i \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   168
  shows "tl (moment (Suc i) s) = moment i s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   169
  using assms unfolding moment_def rev_take
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   170
  by (simp, metis Suc_diff_le diff_Suc_Suc drop_Suc tl_drop)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   171
  
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   172
lemma moment_plus':
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   173
  assumes "Suc i \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   174
  shows "(moment (Suc i) s) = (hd (moment (Suc i) s)) # (moment i s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   175
proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   176
  have "(moment (Suc i) s) \<noteq> []"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   177
  using assms length_moment_le by fastforce 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   178
  hence "(moment (Suc i) s) = (hd (moment (Suc i) s)) #  tl (moment (Suc i) s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   179
    by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   180
  with moment_Suc_tl[OF assms]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   181
  show ?thesis by metis
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   182
qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   183
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   184
lemma moment_plus: 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   185
  "Suc i \<le> length s \<Longrightarrow> moment (Suc i) s = (hd (moment (Suc i) s)) # (moment i s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   186
proof(induct s, simp+)
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   187
  fix a s
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   188
  assume ih: "Suc i \<le> length s \<Longrightarrow> moment (Suc i) s = hd (moment (Suc i) s) # moment i s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   189
    and le_i: "i \<le> length s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   190
  show "moment (Suc i) (a # s) = hd (moment (Suc i) (a # s)) # moment i (a # s)"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   191
  proof(cases "i= length s")
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   192
    case True
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   193
    hence "Suc i = length (a#s)" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   194
    with moment_eq have "moment (Suc i) (a#s) = a#s" by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   195
    moreover have "moment i (a#s) = s"
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   196
    proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   197
      from moment_app [OF le_i, of "[a]"]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   198
      and True show ?thesis by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   199
    qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   200
    ultimately show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   201
  next
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   202
    case False
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   203
    from False and le_i have lti: "i < length s" by arith
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   204
    hence les_i: "Suc i \<le> length s" by arith
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   205
    show ?thesis 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   206
    proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   207
      from moment_app [OF les_i, of "[a]"]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   208
      have "moment (Suc i) (a # s) = moment (Suc i) s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   209
      moreover have "moment i (a#s) = moment i s" 
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   210
      proof -
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   211
        from lti have "i \<le> length s" by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   212
        from moment_app [OF this, of "[a]"] show ?thesis by simp
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   213
      qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   214
      moreover note ih [OF les_i]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   215
      ultimately show ?thesis by auto
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   216
    qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   217
  qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   218
=======
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   219
  with moment_Suc_tl[OF assms]
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   220
  show ?thesis by metis
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   221
>>>>>>> other
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   222
qed
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   223
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   224
end
c495eb16beb6 CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
zhangx
parents:
diff changeset
   225