| author | blanchet | 
| Sun, 25 Apr 2010 15:04:20 +0200 | |
| changeset 36395 | e73923451f6f | 
| parent 34120 | f9920a3ddf50 | 
| child 36960 | 01594f816e3a | 
| permissions | -rw-r--r-- | 
| 14350 | 1  | 
(* Title: HOL/Tools/refute_isar.ML  | 
2  | 
Author: Tjark Weber  | 
|
| 22092 | 3  | 
Copyright 2003-2007  | 
| 14350 | 4  | 
|
| 32855 | 5  | 
Outer syntax commands 'refute' and 'refute_params'.  | 
| 14350 | 6  | 
*)  | 
7  | 
||
| 32855 | 8  | 
structure Refute_Isar: sig end =  | 
| 14350 | 9  | 
struct  | 
10  | 
||
| 32855 | 11  | 
(* argument parsing *)  | 
| 14350 | 12  | 
|
| 32855 | 13  | 
(*optional list of arguments of the form [name1=value1, name2=value2, ...]*)  | 
| 14350 | 14  | 
|
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
33291 
diff
changeset
 | 
15  | 
val scan_parm =  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
33291 
diff
changeset
 | 
16  | 
OuterParse.name  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
33291 
diff
changeset
 | 
17  | 
-- (Scan.optional (OuterParse.$$$ "=" |-- OuterParse.name) "true")  | 
| 32855 | 18  | 
val scan_parms = Scan.optional  | 
19  | 
(OuterParse.$$$ "[" |-- OuterParse.list scan_parm --| OuterParse.$$$ "]") [];  | 
|
| 14350 | 20  | 
|
21  | 
||
| 32855 | 22  | 
(* 'refute' command *)  | 
| 14350 | 23  | 
|
| 32855 | 24  | 
val _ =  | 
25  | 
OuterSyntax.improper_command "refute"  | 
|
26  | 
"try to find a model that refutes a given subgoal" OuterKeyword.diag  | 
|
| 
32857
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
32855 
diff
changeset
 | 
27  | 
(scan_parms -- Scan.optional OuterParse.nat 1 >>  | 
| 
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
32855 
diff
changeset
 | 
28  | 
(fn (parms, i) =>  | 
| 
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
32855 
diff
changeset
 | 
29  | 
Toplevel.keep (fn state =>  | 
| 
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
32855 
diff
changeset
 | 
30  | 
let  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
33291 
diff
changeset
 | 
31  | 
val ctxt = Toplevel.context_of state;  | 
| 33291 | 32  | 
            val {goal = st, ...} = Proof.raw_goal (Toplevel.proof_of state);
 | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
33291 
diff
changeset
 | 
33  | 
in Refute.refute_goal ctxt parms st i end)));  | 
| 14350 | 34  | 
|
35  | 
||
| 32855 | 36  | 
(* 'refute_params' command *)  | 
| 14350 | 37  | 
|
| 32855 | 38  | 
val _ =  | 
39  | 
OuterSyntax.command "refute_params"  | 
|
40  | 
"show/store default parameters for the 'refute' command" OuterKeyword.thy_decl  | 
|
41  | 
(scan_parms >> (fn parms =>  | 
|
42  | 
Toplevel.theory (fn thy =>  | 
|
43  | 
let  | 
|
| 
32857
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
32855 
diff
changeset
 | 
44  | 
val thy' = fold Refute.set_default_param parms thy;  | 
| 32855 | 45  | 
val output =  | 
46  | 
(case Refute.get_default_params thy' of  | 
|
47  | 
[] => "none"  | 
|
| 
32857
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
32855 
diff
changeset
 | 
48  | 
| new_defaults => cat_lines (map (fn (x, y) => x ^ "=" ^ y) new_defaults));  | 
| 32855 | 49  | 
          val _ = writeln ("Default parameters for 'refute':\n" ^ output);
 | 
| 
32857
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
32855 
diff
changeset
 | 
50  | 
in thy' end)));  | 
| 14350 | 51  | 
|
52  | 
end;  | 
|
53  |