3807
|
1 |
(*
|
|
2 |
File: Buffer.ML
|
|
3 |
Author: Stephan Merz
|
|
4 |
Copyright: 1997 University of Munich
|
|
5 |
|
|
6 |
Simple FIFO buffer (theorems and proofs)
|
|
7 |
*)
|
|
8 |
|
|
9 |
(* ---------------------------- Data lemmas ---------------------------- *)
|
|
10 |
|
6255
|
11 |
context List.thy;
|
3807
|
12 |
goal List.thy "xs ~= [] --> tl xs ~= xs";
|
4089
|
13 |
by (auto_tac (claset(), simpset() addsimps [neq_Nil_conv]));
|
3807
|
14 |
qed_spec_mp "tl_not_self";
|
6255
|
15 |
context Buffer.thy;
|
3807
|
16 |
|
6255
|
17 |
Addsimps [tl_not_self];
|
3807
|
18 |
|
|
19 |
(* ---------------------------- Action lemmas ---------------------------- *)
|
|
20 |
|
|
21 |
(* Dequeue is visible *)
|
6255
|
22 |
Goal "|- <Deq ic q oc>_(ic,q,oc) = Deq ic q oc";
|
4089
|
23 |
by (auto_tac (claset(), simpset() addsimps [angle_def,Deq_def]));
|
3807
|
24 |
qed "Deq_visible";
|
|
25 |
|
|
26 |
(* Enabling condition for dequeue -- NOT NEEDED *)
|
5069
|
27 |
Goalw [temp_rewrite Deq_visible]
|
6255
|
28 |
"!!q. basevars (ic,q,oc) ==> |- Enabled (<Deq ic q oc>_(ic,q,oc)) = (q ~= #[])";
|
5525
|
29 |
by (force_tac (claset() addSEs [base_enabled,enabledE], simpset() addsimps [Deq_def]) 1);
|
3807
|
30 |
qed "Deq_enabled";
|
|
31 |
|
|
32 |
(* For the left-to-right implication, we don't need the base variable stuff *)
|
5069
|
33 |
Goalw [temp_rewrite Deq_visible]
|
6255
|
34 |
"|- Enabled (<Deq ic q oc>_(ic,q,oc)) --> (q ~= #[])";
|
4089
|
35 |
by (auto_tac (claset() addSEs [enabledE], simpset() addsimps [Deq_def]));
|
3807
|
36 |
qed "Deq_enabledE";
|