17309
|
1 |
|
|
2 |
(* $Id$ *)
|
|
3 |
|
6255
|
4 |
local
|
17309
|
5 |
fun prover s = prove_goal (the_context ()) s
|
6255
|
6 |
(K [force_tac (claset(), simpset() addsimps [Init_def]) 1])
|
|
7 |
in
|
|
8 |
val const_simps = map (int_rewrite o prover)
|
|
9 |
[ "|- (Init #True) = #True",
|
|
10 |
"|- (Init #False) = #False"]
|
|
11 |
val Init_simps = map (int_rewrite o prover)
|
|
12 |
[ "|- (Init ~F) = (~ Init F)",
|
|
13 |
"|- (Init (P --> Q)) = (Init P --> Init Q)",
|
|
14 |
"|- (Init (P & Q)) = (Init P & Init Q)",
|
|
15 |
"|- (Init (P | Q)) = (Init P | Init Q)",
|
|
16 |
"|- (Init (P = Q)) = ((Init P) = (Init Q))",
|
|
17 |
"|- (Init (!x. F x)) = (!x. (Init F x))",
|
|
18 |
"|- (Init (? x. F x)) = (? x. (Init F x))",
|
|
19 |
"|- (Init (?! x. F x)) = (?! x. (Init F x))"
|
|
20 |
]
|
|
21 |
end;
|
|
22 |
|
|
23 |
Addsimps const_simps;
|
|
24 |
|
|
25 |
Goal "|- (Init $P) = (Init P)";
|
|
26 |
by (force_tac (claset(), simpset() addsimps [Init_def,fw_act_def,fw_stp_def]) 1);
|
|
27 |
qed "Init_stp_act";
|
|
28 |
val Init_simps = (int_rewrite Init_stp_act)::Init_simps;
|
|
29 |
bind_thm("Init_stp_act_rev", symmetric(int_rewrite Init_stp_act));
|
|
30 |
|
|
31 |
Goal "|- (Init F) = F";
|
|
32 |
by (force_tac (claset(), simpset() addsimps [Init_def,fw_temp_def]) 1);
|
|
33 |
qed "Init_temp";
|
|
34 |
val Init_simps = (int_rewrite Init_temp)::Init_simps;
|
|
35 |
|
|
36 |
(* Trivial instances of the definitions that avoid introducing lambda expressions. *)
|
|
37 |
Goalw [Init_def,fw_stp_def] "(sigma |= Init P) = P (st1 sigma)";
|
|
38 |
by (rtac refl 1);
|
|
39 |
qed "Init_stp";
|
|
40 |
|
|
41 |
Goalw [Init_def,fw_act_def] "(sigma |= Init A) = A (st1 sigma, st2 sigma)";
|
|
42 |
by (rtac refl 1);
|
|
43 |
qed "Init_act";
|
|
44 |
|
|
45 |
val Init_defs = [Init_stp, Init_act, int_use Init_temp];
|