3807
|
1 |
(* Title: TLA/cladata.ML
|
|
2 |
Author: Stephan Merz (mostly stolen from Isabelle/HOL)
|
|
3 |
|
|
4 |
Setting up the classical reasoner for TLA.
|
|
5 |
|
|
6 |
The classical prover for TLA uses a different hyp_subst_tac that substitutes
|
|
7 |
somewhat more liberally for state variables. Unfortunately, this requires
|
|
8 |
either generating a new prover or redefining the basic proof tactics.
|
|
9 |
We take the latter approach, because otherwise there would be a type conflict
|
|
10 |
between standard HOL and TLA classical sets, and we would have to redefine
|
|
11 |
even more things (e.g., blast_tac), and try to keep track of which rules
|
|
12 |
have been active in setting up a new default claset.
|
|
13 |
|
|
14 |
*)
|
|
15 |
|
|
16 |
|
|
17 |
(* Generate a different hyp_subst_tac
|
|
18 |
that substitutes for x(s) if s is a bound variable of "world" type.
|
|
19 |
This is useful to solve equations that contain state variables.
|
|
20 |
*)
|
|
21 |
|
|
22 |
use "hypsubst.ML"; (* local version! *)
|
|
23 |
|
|
24 |
structure ActHypsubst_Data =
|
|
25 |
struct
|
|
26 |
structure Simplifier = Simplifier
|
|
27 |
(*Take apart an equality judgement; otherwise raise Match!*)
|
|
28 |
fun dest_eq (Const("Trueprop",_) $ (Const("op =",_) $ t $ u)) = (t,u);
|
|
29 |
val eq_reflection = eq_reflection
|
|
30 |
val imp_intr = impI
|
|
31 |
val rev_mp = rev_mp
|
|
32 |
val subst = subst
|
|
33 |
val sym = sym
|
|
34 |
end;
|
|
35 |
|
|
36 |
structure ActHypsubst = ActHypsubstFun(ActHypsubst_Data);
|
|
37 |
open ActHypsubst;
|
|
38 |
|
|
39 |
|
|
40 |
(**
|
|
41 |
Define the basic classical set and clasimpset for the action part of TLA.
|
|
42 |
Add the new hyp_subst_tac to the wrapper (also for the default claset).
|
|
43 |
**)
|
|
44 |
|
4651
|
45 |
val action_css = (HOL_cs addSIs [actionI,intI] addSEs [exE_prop]
|
|
46 |
addDs [actionD,intD]
|
|
47 |
addSaltern ("action_hyp_subst_tac", action_hyp_subst_tac),
|
|
48 |
simpset());
|
|
49 |
val action_cs = op addss action_css;
|
3807
|
50 |
|
|
51 |
|
|
52 |
AddSIs [actionI,intI];
|
|
53 |
AddDs [actionD,intD];
|
|
54 |
|
|
55 |
|
|
56 |
|
4658
|
57 |
|