Moment.thy
author zhangx
Thu, 28 Jan 2016 21:14:17 +0800
changeset 90 ed938e2246b9
parent 75 2aa37de77f31
child 100 3d2b59f15f26
permissions -rw-r--r--
Retrofiting of: CpsG.thy (the parallel copy of PIPBasics.thy), ExtGG.thy (The paralell copy of Implemenation.thy), PrioG.thy (The paralell copy of Correctness.thy) has completed. The next step is to overwite original copies with the paralell ones.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
theory Moment
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
imports Main
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
69
1dc801552dfd simplified Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
     5
definition moment :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
1dc801552dfd simplified Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
     6
where "moment n s = rev (take n (rev s))"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 35
diff changeset
     7
70
92ca2410b3d9 further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 69
diff changeset
     8
value "moment 3 [0, 1, 2, 3, 4, 5, 6, 7, 8, 9::int]"
92ca2410b3d9 further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 69
diff changeset
     9
value "moment 2 [5, 4, 3, 2, 1, 0::int]"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
lemma moment_app [simp]:
69
1dc801552dfd simplified Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
    12
  assumes ile: "i \<le> length s"
70
92ca2410b3d9 further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 69
diff changeset
    13
  shows "moment i (s' @ s) = moment i s"
69
1dc801552dfd simplified Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
    14
using assms unfolding moment_def by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
70
92ca2410b3d9 further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 69
diff changeset
    16
lemma moment_eq [simp]: "moment (length s) (s' @ s) = s"
69
1dc801552dfd simplified Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
    17
  unfolding moment_def by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
lemma moment_ge [simp]: "length s \<le> n \<Longrightarrow> moment n s = s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
  by (unfold moment_def, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
lemma moment_zero [simp]: "moment 0 s = []"
69
1dc801552dfd simplified Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
    23
  by (simp add:moment_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
74
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    25
lemma least_idx:
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    26
  assumes "Q (i::nat)"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    27
  obtains j where "j \<le> i" "Q j" "\<forall> k < j. \<not> Q k"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    28
  using assms
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    29
  by (metis ex_least_nat_le le0 not_less0) 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    30
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    31
lemma duration_idx:
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    32
  assumes "\<not> Q (i::nat)"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    33
  and "Q j"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    34
  and "i \<le> j"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    35
  obtains k where "i \<le> k" "k < j" "\<not> Q k" "\<forall> i'. k < i' \<and> i' \<le> j \<longrightarrow> Q i'" 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    36
proof -
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    37
  let ?Q = "\<lambda> t. t \<le> j \<and> \<not> Q (j - t)"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    38
  have "?Q (j - i)" using assms by (simp add: assms(1)) 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    39
  from least_idx [of ?Q, OF this]
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    40
  obtain l
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    41
  where h: "l \<le> j - i" "\<not> Q (j - l)" "\<forall>k<l. \<not> (k \<le> j \<and> \<not> Q (j - k))"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    42
    by metis
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    43
  let ?k = "j - l"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    44
  have "i \<le> ?k" using assms(3) h(1) by linarith 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    45
  moreover have "?k < j" by (metis assms(2) diff_le_self h(2) le_neq_implies_less) 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    46
  moreover have "\<not> Q ?k" by (simp add: h(2)) 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    47
  moreover have "\<forall> i'. ?k < i' \<and> i' \<le> j \<longrightarrow> Q i'"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    48
      by (metis diff_diff_cancel diff_le_self diff_less_mono2 h(3) 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    49
              less_imp_diff_less not_less) 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    50
  ultimately show ?thesis using that by metis
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
74
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    53
lemma p_split_gen: 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    54
  assumes "Q s"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    55
  and "\<not> Q (moment k s)"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    56
  shows "(\<exists> i. i < length s \<and> k \<le> i \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    57
proof(cases "k \<le> length s")
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    58
  case True
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    59
  let ?Q = "\<lambda> t. Q (moment t s)"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    60
  have "?Q (length s)" using assms(1) by simp
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    61
  from duration_idx[of ?Q, OF assms(2) this True]
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    62
  obtain i where h: "k \<le> i" "i < length s" "\<not> Q (moment i s)"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    63
    "\<forall>i'. i < i' \<and> i' \<le> length s \<longrightarrow> Q (moment i' s)" by metis
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    64
  moreover have "(\<forall> i' > i. Q (moment i' s))" using h(4) assms(1) not_less
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    65
    by fastforce
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    66
  ultimately show ?thesis by metis
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    67
qed (insert assms, auto)
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    68
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
lemma p_split: 
74
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    70
  assumes qs: "Q s"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    71
  and nq: "\<not> Q []"
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    72
  shows "(\<exists> i. i < length s \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
  from nq have "\<not> Q (moment 0 s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
  from p_split_gen [of Q s 0, OF qs this]
74
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    76
  show ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
71
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    79
lemma moment_Suc_tl:
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    80
  assumes "Suc i \<le> length s"
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    81
  shows "tl (moment (Suc i) s) = moment i s"
74
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    82
  using assms 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    83
  by (simp add:moment_def rev_take, 
83ba2d8c859a Moment.thy further simplified.
zhangx
parents: 73
diff changeset
    84
      metis Suc_diff_le diff_Suc_Suc drop_Suc tl_drop)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 72
diff changeset
    85
  
71
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    86
lemma moment_plus:
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    87
  assumes "Suc i \<le> length s"
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    88
  shows "(moment (Suc i) s) = (hd (moment (Suc i) s)) # (moment i s)"
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    89
proof -
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 72
diff changeset
    90
  have "(moment (Suc i) s) \<noteq> []" using assms 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 72
diff changeset
    91
    by (simp add:moment_def rev_take)
71
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    92
  hence "(moment (Suc i) s) = (hd (moment (Suc i) s)) #  tl (moment (Suc i) s)"
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    93
    by auto
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    94
  with moment_Suc_tl[OF assms]
04caf0ccb3ae some small change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 70
diff changeset
    95
  show ?thesis by metis
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 0
diff changeset
    98
end
70
92ca2410b3d9 further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 69
diff changeset
    99