-
Notifications
You must be signed in to change notification settings - Fork 0
/
Common_Proof.thy
484 lines (412 loc) · 15.6 KB
/
Common_Proof.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
section "Common Proofs"
theory Common_Proof
imports Common Idle_Proof Current_Proof
begin
lemma reverseN_drop: "reverseN n xs acc = drop (length xs - n) (rev xs) @ acc"
unfolding reverseN_def using rev_take by blast
lemma reverseN_step: "xs \<noteq> [] \<Longrightarrow> reverseN n (tl xs) (hd xs # acc) = reverseN (Suc n) xs acc"
by (simp add: take_Suc)
lemma reverseN_finish: "reverseN n [] acc = acc"
by (simp)
lemma reverseN_tl_hd: "0 < n \<Longrightarrow> xs \<noteq> [] \<Longrightarrow> reverseN n xs ys = reverseN (n - (Suc 0)) (tl xs) (hd xs #ys)"
by (simp add: reverseN_step del: reverseN_def)
lemma reverseN_nth: "n < length xs \<Longrightarrow> x = xs ! n \<Longrightarrow> x # reverseN n xs ys = reverseN (Suc n) xs ys"
by (simp add: take_Suc_conv_app_nth)
lemma step_list [simp]: "invar common \<Longrightarrow> list (step common) = list common"
proof(induction common rule: step_state.induct)
case (1 idle)
then show ?case by auto
next
case (2 current aux new moved)
then show ?case
proof(cases current)
case (Current extra added old remained)
with 2 have aux_not_empty: "aux \<noteq> []"
by auto
from 2 Current show ?thesis
proof(cases "remained \<le> Suc moved")
case True
with 2 Current have "remained - length new = 1"
by auto
with True Current 2 aux_not_empty show ?thesis
by(auto simp: )
next
case False
with Current show ?thesis
by(auto simp: aux_not_empty reverseN_step Suc_diff_Suc simp del: reverseN_def)
qed
qed
qed
lemma step_list_current [simp]: "invar common \<Longrightarrow> list_current (step common) = list_current common"
by(cases common)(auto split: current.splits)
lemma push_list [simp]: "list (push x common) = x # list common"
proof(induction x common rule: push.induct)
case (1 x stack stackSize)
then show ?case
by auto
next
case (2 x current aux new moved)
then show ?case
by(induction x current rule: Current.push.induct) auto
qed
lemma invar_step: "invar (common :: 'a state) \<Longrightarrow> invar (step common)"
proof(induction "common" rule: invar_state.induct)
case (1 idle)
then show ?case
by auto
next
case (2 current aux new moved)
then show ?case
proof(cases current)
case (Current extra added old remained)
then show ?thesis
proof(cases "aux = []")
case True
with 2 Current show ?thesis by auto
next
case False
note AUX_NOT_EMPTY = False
then show ?thesis
proof(cases "remained \<le> Suc (length new)")
case True
with 2 Current False
have "take (Suc (length new)) (Stack.list old) = take (size old) (hd aux # new)"
by(auto simp: le_Suc_eq take_Cons')
with 2 Current True show ?thesis
by auto
next
case False
with 2 Current AUX_NOT_EMPTY show ?thesis
by(auto simp: reverseN_step Suc_diff_Suc simp del: reverseN_def)
qed
qed
qed
qed
lemma invar_push: "invar common \<Longrightarrow> invar (push x common)"
proof(induction x common rule: push.induct)
case (1 x current stack stackSize)
then show ?case
proof(induction x current rule: Current.push.induct)
case (1 x extra added old remained)
then show ?case
proof(induction x stack rule: Stack.push.induct)
case (1 x left right)
then show ?case by auto
qed
qed
next
case (2 x current aux new moved)
then show ?case
proof(induction x current rule: Current.push.induct)
case (1 x extra added old remained)
then show ?case by auto
qed
qed
lemma invar_pop: "\<lbrakk>
0 < size common;
invar common;
pop common = (x, common')
\<rbrakk> \<Longrightarrow> invar common'"
proof(induction common arbitrary: x rule: pop.induct)
case (1 current idle)
then obtain idle' where idle: "Idle.pop idle = (x, idle')"
by(auto split: prod.splits)
obtain current' where current: "drop_first current = current'"
by auto
from 1 current idle show ?case
using Idle_Proof.size_pop[of idle x idle', symmetric]
size_new_pop[of current]
size_pop_sub[of current _ current']
by(auto simp: Idle_Proof.invar_pop invar_size_pop eq_snd_iff take_tl size_not_empty)
next
case (2 current aux new moved)
then show ?case
proof(induction current rule: Current.pop.induct)
case (1 added old remained)
then show ?case
proof(cases "remained - Suc 0 \<le> length new")
case True
with 1 have [simp]:
"0 < size old"
"Stack.list old \<noteq> []"
"aux \<noteq> []"
"length new = remained - Suc 0"
by(auto simp: Stack_Proof.size_not_empty Stack_Proof.list_not_empty)
then have [simp]: "Suc 0 \<le> size old"
by linarith
from 1 have "0 < remained"
by auto
then have "take remained (Stack.list old)
= hd (Stack.list old) # take (remained - Suc 0) (tl (Stack.list old))"
by (metis Suc_pred \<open>Stack.list old \<noteq> []\<close> list.collapse take_Suc_Cons)
with 1 True show ?thesis
using Stack_Proof.pop_list[of old]
by(auto simp: Stack_Proof.size_not_empty)
next
case False
with 1 have "remained - Suc 0 \<le> length aux + length new" by auto
with 1 False show ?thesis
using Stack_Proof.pop_list[of old]
apply(auto simp: Suc_diff_Suc take_tl Stack_Proof.size_not_empty tl_append_if)
by (simp add: Suc_diff_le rev_take tl_drop_2 tl_take)
qed
next
case (2 x xs added old remained)
then show ?case by auto
qed
qed
lemma push_list_current [simp]: "list_current (push x left) = x # list_current left"
by(induction x left rule: push.induct) auto
lemma pop_list [simp]: "invar common \<Longrightarrow> 0 < size common \<Longrightarrow> pop common = (x, common') \<Longrightarrow>
x # list common' = list common"
proof(induction common arbitrary: x rule: pop.induct)
case 1
then show ?case
by(auto simp: size_not_empty split: prod.splits)
next
case (2 current aux new moved)
then show ?case
proof(induction current rule: Current.pop.induct)
case (1 added old remained)
then show ?case
proof(cases "remained - Suc 0 \<le> length new")
case True
from 1 True have [simp]:
"aux \<noteq> []" "0 < remained"
"Stack.list old \<noteq> []" "remained - length new = 1"
by(auto simp: Stack_Proof.size_not_empty Stack_Proof.list_not_empty)
then have "take remained (Stack.list old) = hd aux # take (size old - Suc 0) new
\<Longrightarrow> Stack.first old = hd aux"
by (metis first_hd hd_take list.sel(1))
with 1 True take_hd[of aux] show ?thesis
by(auto simp: Suc_leI)
next
case False
then show ?thesis
proof(cases "remained - length new = length aux")
case True
then have length_minus_1: "remained - Suc (length new) = length aux - 1"
by simp
from 1 have not_empty: "0 < remained" "0 < size old" "aux \<noteq> []" "\<not> is_empty old"
by(auto simp: Stack_Proof.size_not_empty)
from 1 True not_empty have "take 1 (Stack.list old) = take 1 (rev aux)"
using take_1[of
remained
"size old"
"Stack.list old"
"(rev aux) @ take (size old + length new - remained) new"
]
by(simp)
then have "[last aux] = [Stack.first old]"
using take_last first_take not_empty
by fastforce
then have "last aux = Stack.first old"
by auto
with 1 True False show ?thesis
using not_empty last_drop_rev[of aux]
by(auto simp: reverseN_drop length_minus_1 simp del: reverseN_def)
next
case False
with 1 have a: "take (remained - length new) aux \<noteq> []"
by auto
from 1 False have b: "\<not> is_empty old"
by(auto simp: Stack_Proof.size_not_empty)
from 1 have c: "remained - Suc (length new) < length aux"
by auto
from 1 have not_empty: "0 < remained" "0 < size old" "0 < remained - length new" "0 < length aux"
by auto
with False have "
take remained (Stack.list old) = take (size old) (reverseN (remained - length new) aux new)
\<Longrightarrow> take (Suc 0) (Stack.list old) = take (Suc 0) (rev (take (remained - length new) aux))"
using take_1[of
remained
"size old"
"Stack.list old"
" (reverseN (remained - length new) aux new)"
]
by(auto simp: not_empty Suc_le_eq)
with 1 False have "take 1 (Stack.list old) = take 1 (rev (take (remained - length new) aux))"
by auto
then have d: "[Stack.first old] = [last (take (remained - length new) aux)]"
using take_last first_take a b
by metis
have "last (take (remained - length new) aux) # rev (take (remained - Suc (length new)) aux)
= rev (take (remained - length new) aux)"
using Suc_diff_Suc c not_empty
by (metis a drop_drop last_drop_rev plus_1_eq_Suc rev_take zero_less_diff)
with 1(1) 1(3) False not_empty d show ?thesis
by(cases "remained - length new = 1") (auto)
qed
qed
next
case 2
then show ?case by auto
qed
qed
lemma pop_list_current: "invar common \<Longrightarrow> 0 < size common \<Longrightarrow> pop common = (x, common')
\<Longrightarrow> x # list_current common' = list_current common"
proof(induction common arbitrary: x rule: pop.induct)
case (1 current idle)
then show ?case
proof(induction idle rule: Idle.pop.induct)
case (1 stack stackSize)
then show ?case
proof(induction current rule: Current.pop.induct)
case (1 added old remained)
then have "Stack.first old = Stack.first stack"
using take_first[of old stack]
by auto
with 1 show ?case
by(auto simp: Stack_Proof.size_not_empty Stack_Proof.list_not_empty)
next
case (2 x xs added old remained)
then have "0 < size stack"
by auto
with Stack_Proof.size_not_empty Stack_Proof.list_not_empty
have not_empty: "\<not> is_empty stack" "Stack.list stack \<noteq> []"
by auto
with 2 have "hd (Stack.list stack) = x"
using take_hd'[of "Stack.list stack" x "xs @ Stack.list old"]
by auto
with 2 show ?case
using first_list[of stack] not_empty
by auto
qed
qed
next
case (2 current)
then show ?case
proof(induction current rule: Current.pop.induct)
case (1 added old remained)
then have "\<not> is_empty old"
by(auto simp: Stack_Proof.size_not_empty)
with 1 show ?case
using first_pop
by(auto simp: Stack_Proof.list_not_empty)
next
case 2
then show ?case by auto
qed
qed
lemma list_current_size [simp]:
"\<lbrakk>0 < size common; list_current common = []; invar common\<rbrakk> \<Longrightarrow> False"
proof(induction common rule: invar_state.induct)
case 1
then show ?case
using list_size by auto
next
case (2 current)
then have "invar current"
"Current.list current = []"
"0 < size current"
by(auto split: current.splits)
then show ?case using list_size by auto
qed
lemma list_size [simp]: "\<lbrakk>0 < size common; list common = []; invar common\<rbrakk> \<Longrightarrow> False"
proof(induction common rule: invar_state.induct)
case 1
then show ?case
using list_size Idle_Proof.size_empty
by auto
next
case (2 current aux new moved)
then have "invar current"
"Current.list current = []"
"0 < size current"
by(auto split: current.splits)
then show ?case using list_size by auto
qed
lemma size_empty: "invar (common :: 'a state) \<Longrightarrow> size common = 0 \<Longrightarrow> is_empty common"
proof(induction common rule: is_empty_state.induct)
case 1
then show ?case
by(auto simp: min_def size_empty size_new_empty split: if_splits)
next
case (2 current)
then have "invar current"
by(auto split: current.splits)
with 2 show ?case
by(auto simp: min_def size_empty size_new_empty split: if_splits)
qed
lemma step_size [simp]: "invar (common :: 'a state) \<Longrightarrow> size (step common) = size common"
proof(induction common rule: step_state.induct)
case 1
then show ?case by auto
next
case 2
then show ?case
by(auto simp: min_def split: current.splits)
qed
lemma step_size_new [simp]: "invar (common :: 'a state)
\<Longrightarrow> size_new (step common) = size_new common"
proof(induction common rule: step_state.induct)
case (1 current idle)
then show ?case by auto
next
case (2 current aux new moved)
then show ?case by(auto split: current.splits)
qed
lemma remaining_steps_step [simp]: "\<lbrakk>invar (common :: 'a state); remaining_steps common > 0\<rbrakk>
\<Longrightarrow> Suc (remaining_steps (step common)) = remaining_steps common"
by(induction common)(auto split: current.splits)
lemma remaining_steps_step_sub [simp]: "\<lbrakk>invar (common :: 'a state)\<rbrakk>
\<Longrightarrow> remaining_steps (step common) = remaining_steps common - 1"
by(induction common)(auto split: current.splits)
lemma remaining_steps_step_0 [simp]: "\<lbrakk>invar (common :: 'a state); remaining_steps common = 0\<rbrakk>
\<Longrightarrow> remaining_steps (step common) = 0"
by(induction common)(auto split: current.splits)
lemma remaining_steps_push [simp]: "invar common
\<Longrightarrow> remaining_steps (push x common) = remaining_steps common"
by(induction x common rule: Common.push.induct)(auto split: current.splits)
lemma remaining_steps_pop: "\<lbrakk>invar common; 0 < size common; pop common = (x, common')\<rbrakk>
\<Longrightarrow> remaining_steps common' \<le> remaining_steps common"
proof(induction common rule: pop.induct)
case (1 current idle)
then show ?case
proof(induction idle rule: Idle.pop.induct)
case 1
then show ?case
by(induction current rule: Current.pop.induct) auto
qed
next
case (2 current aux new moved)
then show ?case
by(induction current rule: Current.pop.induct) auto
qed
lemma size_push [simp]: "invar common \<Longrightarrow> size (push x common) = Suc (size common)"
by(induction x common rule: push.induct) (auto split: current.splits)
lemma size_new_push [simp]: "invar common \<Longrightarrow> size_new (push x common) = Suc (size_new common)"
by(induction x common rule: Common.push.induct) (auto split: current.splits)
lemma size_pop [simp]: "\<lbrakk>invar common; 0 < size common; pop common = (x, common')\<rbrakk>
\<Longrightarrow> Suc (size common') = size common"
proof(induction common rule: Common.pop.induct)
case (1 current idle)
then show ?case
using size_drop_first_sub[of current] Idle_Proof.size_pop_sub[of idle]
by(auto simp: size_not_empty split: prod.splits)
next
case (2 current aux new moved)
then show ?case
by(induction current rule: Current.pop.induct) auto
qed
lemma size_new_pop [simp]: "\<lbrakk>invar common; 0 < size_new common; pop common = (x, common')\<rbrakk>
\<Longrightarrow> Suc (size_new common') = size_new common"
proof(induction common rule: Common.pop.induct)
case (1 current idle)
then show ?case
using size_new_pop[of current]
by(auto split: prod.splits)
next
case (2 current aux new moved)
then show ?case
proof(induction current rule: Current.pop.induct)
case (1 added old remained)
then show ?case by auto
next
case (2 x xs added old remained)
then show ?case by auto
qed
qed
lemma size_size_new: "\<lbrakk>invar (common :: 'a state); 0 < size common\<rbrakk> \<Longrightarrow> 0 < size_new common"
by(cases common) auto
end