3210
|
1 |
(* Title: HOL/Modelcheck/MCSyn.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Olaf Mueller, Jan Philipps, Robert Sandner
|
|
4 |
Copyright 1997 TU Muenchen
|
|
5 |
*)
|
|
6 |
|
|
7 |
fun mc_tac i state =
|
|
8 |
let val sign = #sign (rep_thm state)
|
|
9 |
in
|
|
10 |
case drop(i-1,prems_of state) of
|
|
11 |
[] => Sequence.null |
|
|
12 |
subgoal::_ =>
|
|
13 |
let val concl = Logic.strip_imp_concl subgoal;
|
|
14 |
val OraAss = invoke_oracle(MCSyn.thy,sign,MCOracleExn concl);
|
|
15 |
in
|
|
16 |
((cut_facts_tac [OraAss] i) THEN (atac i)) state
|
|
17 |
end
|
|
18 |
end;
|
|
19 |
|
|
20 |
|
|
21 |
goal Prod.thy "(? x. P x) = (? a b. P(a,b))";
|
|
22 |
auto();
|
|
23 |
by (split_all_tac 1);
|
|
24 |
auto();
|
|
25 |
qed "split_paired_Ex";
|
|
26 |
|
|
27 |
|
|
28 |
goalw thy [split_def] "(f::'a*'b=>'c) = (%(x, y). f (x, y))";
|
|
29 |
br ext 1;
|
|
30 |
by (stac (surjective_pairing RS sym) 1);
|
|
31 |
br refl 1;
|
|
32 |
qed "pair_eta_expand";
|
|
33 |
|
|
34 |
local
|
|
35 |
val lhss = [cterm_of (sign_of thy) (read "f::'a*'b=>'c")];
|
|
36 |
val rew = mk_meta_eq pair_eta_expand;
|
|
37 |
|
|
38 |
fun proc _ (Abs _) = Some rew
|
|
39 |
| proc _ _ = None;
|
|
40 |
in
|
|
41 |
val pair_eta_expand_proc = Simplifier.mk_simproc "pair_eta_expand" lhss proc;
|
|
42 |
end;
|
|
43 |
|
|
44 |
|
|
45 |
val MC_ss = (!simpset addsimprocs [pair_eta_expand_proc])
|
|
46 |
addsimps [split_paired_Ex,Let_def];
|