8177
|
1 |
(* Title: HOL/IMPP/Misc.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: David von Oheimb
|
|
4 |
Copyright 1999 TUM
|
|
5 |
*)
|
|
6 |
|
|
7 |
section "state access";
|
|
8 |
|
|
9 |
Goalw [getlocs_def] "getlocs (st g l) = l";
|
|
10 |
by (Simp_tac 1);
|
|
11 |
qed "getlocs_def2";
|
|
12 |
|
|
13 |
Goalw [update_def] "s[Loc Y::=s<Y>] = s";
|
|
14 |
by (induct_tac "s" 1);
|
|
15 |
by (simp_tac (simpset() addsimps [getlocs_def2]) 1);
|
|
16 |
qed "update_Loc_idem2";
|
|
17 |
Addsimps[update_Loc_idem2];
|
|
18 |
|
|
19 |
Goalw [update_def] "s[X::=x][X::=y] = s[X::=y]";
|
|
20 |
by (induct_tac "X" 1);
|
|
21 |
by Auto_tac;
|
|
22 |
by (ALLGOALS (induct_tac "s"));
|
|
23 |
by Auto_tac;
|
|
24 |
by (ALLGOALS (rtac ext));
|
|
25 |
by Auto_tac;
|
|
26 |
qed "update_overwrt";
|
|
27 |
Addsimps[update_overwrt];
|
|
28 |
|
|
29 |
Goalw [update_def]"(s[Loc Y::=k])<Y'> = (if Y=Y' then k else s<Y'>)";
|
|
30 |
by (induct_tac "s" 1);
|
|
31 |
by (simp_tac (simpset() addsimps [getlocs_def2]) 1);
|
|
32 |
qed "getlocs_Loc_update";
|
|
33 |
Addsimps[getlocs_Loc_update];
|
|
34 |
|
|
35 |
Goalw [update_def] "getlocs (s[Glb Y::=k]) = getlocs s";
|
|
36 |
by (induct_tac "s" 1);
|
|
37 |
by (simp_tac (simpset() addsimps [getlocs_def2]) 1);
|
|
38 |
qed "getlocs_Glb_update";
|
|
39 |
Addsimps[getlocs_Glb_update];
|
|
40 |
|
|
41 |
Goalw [setlocs_def] "getlocs (setlocs s l) = l";
|
|
42 |
by (induct_tac "s" 1);
|
|
43 |
by Auto_tac;
|
|
44 |
by (simp_tac (simpset() addsimps [getlocs_def2]) 1);
|
|
45 |
qed "getlocs_setlocs";
|
|
46 |
Addsimps[getlocs_setlocs];
|
|
47 |
|
|
48 |
Goal "getlocs (setlocs s (getlocs s')[Y::=k]) = getlocs (s'[Y::=k])";
|
|
49 |
by (induct_tac "Y" 1);
|
10962
|
50 |
by (rtac ext 2);
|
8177
|
51 |
by Auto_tac;
|
|
52 |
qed "getlocs_setlocs_lemma";
|
|
53 |
|
|
54 |
(*unused*)
|
|
55 |
Goalw [hoare_valids_def]
|
|
56 |
"!v. G|={%Z s. P Z (s[Loc Y::=v]) & s<Y> = a (s[Loc Y::=v])}. \
|
|
57 |
\ c .{%Z s. Q Z (s[Loc Y::=v])} ==> G|={P}. LOCAL Y:=a IN c .{Q}";
|
|
58 |
by (full_simp_tac (simpset() addsimps [triple_valid_def2]) 1);
|
|
59 |
by (Clarsimp_tac 1);
|
|
60 |
by (dres_inst_tac [("x","s<Y>")] spec 1);
|
|
61 |
by (smp_tac 1 1);
|
10962
|
62 |
by (dtac spec 1);
|
8177
|
63 |
by (dres_inst_tac [("x","s[Loc Y::=a s]")] spec 1);
|
|
64 |
by (Full_simp_tac 1);
|
|
65 |
by (mp_tac 1);
|
|
66 |
by (smp_tac 1 1);
|
|
67 |
by (Simp_tac 1);
|
|
68 |
qed "classic_Local_valid";
|
|
69 |
|
|
70 |
Goal "!v. G|-{%Z s. P Z (s[Loc Y::=v]) & s<Y> = a (s[Loc Y::=v])}. \
|
|
71 |
\ c .{%Z s. Q Z (s[Loc Y::=v])} ==> G|-{P}. LOCAL Y:=a IN c .{Q}";
|
10962
|
72 |
by (rtac export_s 1);
|
8177
|
73 |
(*variant 1*)
|
10962
|
74 |
by (rtac (hoare_derivs.Local RS conseq1) 1);
|
|
75 |
by (etac spec 1);
|
8177
|
76 |
by (Force_tac 1);
|
|
77 |
(*variant 2
|
|
78 |
by (res_inst_tac [("P'","%Z s. s' = s & P Z (s[Loc Y::=a s][Loc Y::=s'<Y>]) & (s[Loc Y::=a s])<Y> = a (s[Loc Y::=a s][Loc Y::=s'<Y>])")] conseq1 1);
|
|
79 |
by (Clarsimp_tac 2);
|
10962
|
80 |
by (rtac hoare_derivs.Local 1);
|
|
81 |
by (etac spec 1);
|
8177
|
82 |
*)
|
|
83 |
qed "classic_Local";
|
|
84 |
|
|
85 |
Goal "[| Y~=Y'; G|-{P}. c .{%Z s. s<Y'>=d} |] ==> \
|
|
86 |
\ G|-{%Z s. P Z (s[Loc Y::=a s])}. LOCAL Y:=a IN c .{%Z s. s<Y'>=d}";
|
10962
|
87 |
by (rtac classic_Local 1);
|
8177
|
88 |
by (ALLGOALS Clarsimp_tac);
|
10962
|
89 |
by (etac conseq12 1);
|
8177
|
90 |
by (Clarsimp_tac 1);
|
10962
|
91 |
by (dtac sym 1);
|
8177
|
92 |
by (Asm_full_simp_tac 1);
|
|
93 |
qed "classic_Local_indep";
|
|
94 |
|
|
95 |
Goal "[| Y~=Y'; G|-{P}. c .{%Z s. s<Y'>=d} |] ==> \
|
|
96 |
\ G|-{%Z s. P Z (s[Loc Y::=a s])}. LOCAL Y:=a IN c .{%Z s. s<Y'>=d}";
|
10962
|
97 |
by (rtac export_s 1);
|
|
98 |
by (rtac hoare_derivs.Local 1);
|
8177
|
99 |
by (Clarsimp_tac 1);
|
|
100 |
qed "Local_indep";
|
|
101 |
|
|
102 |
Goal "[| Y'~=Y; G|-{P}. c .{%Z s. s<Y'>=d} |] ==> \
|
|
103 |
\ G|-{%Z s. P Z (s[Loc Y::=a s])}. LOCAL Y:=a IN c .{%Z s. s<Y'>=d}";
|
10962
|
104 |
by (rtac weak_Local 1);
|
8177
|
105 |
by (ALLGOALS Clarsimp_tac);
|
|
106 |
qed "weak_Local_indep";
|
|
107 |
|
|
108 |
|
|
109 |
Goal "G|-{%Z s. Z = s<Y>}. LOCAL Y:=a IN c .{%Z s. Z = s<Y>}";
|
10962
|
110 |
by (rtac export_s 1);
|
8177
|
111 |
by (res_inst_tac [("P'","%Z s. s'=s & True"), ("Q'","%Z s. s'<Y> = s<Y>")] (conseq12) 1);
|
|
112 |
by (Clarsimp_tac 2);
|
10962
|
113 |
by (rtac hoare_derivs.Local 1);
|
8177
|
114 |
by (Clarsimp_tac 1);
|
10962
|
115 |
by (rtac trueI 1);
|
8177
|
116 |
qed "export_Local_invariant";
|
|
117 |
|
|
118 |
Goal "G|-{%Z s. Z = s<Y>}. LOCAL Y:=a IN c .{%Z s. Z = s<Y>}";
|
10962
|
119 |
by (rtac classic_Local 1);
|
8177
|
120 |
by (Clarsimp_tac 1);
|
10962
|
121 |
by (rtac (trueI RS conseq12) 1);
|
8177
|
122 |
by (Clarsimp_tac 1);
|
|
123 |
qed "classic_Local_invariant";
|
|
124 |
|
|
125 |
Goal "G|-{P}. BODY pn .{%Z s. Q Z (setlocs s (getlocs s')[X::=s<Res>])} ==> \
|
|
126 |
\ G|-{%Z s. s'=s & I Z (getlocs (s[X::=k Z])) & P Z (setlocs s newlocs[Loc Arg::=a s])}. \
|
|
127 |
\ X:=CALL pn (a) .{%Z s. I Z (getlocs (s[X::=k Z])) & Q Z s}";
|
|
128 |
by (res_inst_tac [("s'1","s'"),("Q'","%Z s. I Z (getlocs (s[X::=k Z])) = I Z (getlocs (s'[X::=k Z])) & Q Z s")] (hoare_derivs.Call RS conseq12) 1);
|
|
129 |
by (asm_simp_tac (simpset() addsimps [getlocs_setlocs_lemma]) 1);
|
|
130 |
by (Force_tac 1);
|
|
131 |
qed "Call_invariant";
|
|
132 |
|
|
133 |
|