| 
14350
 | 
     1  | 
(*  Title:      HOL/Tools/refute_isar.ML
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Tjark Weber
  | 
| 
22092
 | 
     4  | 
    Copyright   2003-2007
  | 
| 
14350
 | 
     5  | 
  | 
| 
 | 
     6  | 
Adds the 'refute' and 'refute_params' commands to Isabelle/Isar's outer
  | 
| 
 | 
     7  | 
syntax.
  | 
| 
 | 
     8  | 
*)
  | 
| 
 | 
     9  | 
  | 
| 
24867
 | 
    10  | 
structure RefuteIsar: sig end =
  | 
| 
14350
 | 
    11  | 
struct
  | 
| 
 | 
    12  | 
  | 
| 
 | 
    13  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    14  | 
(* common functions for argument parsing/scanning                            *)
  | 
| 
 | 
    15  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    16  | 
  | 
| 
 | 
    17  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    18  | 
(* both 'refute' and 'refute_params' take an optional list of arguments of   *)
  | 
| 
 | 
    19  | 
(* the form [name1=value1, name2=value2, ...]                                *)
  | 
| 
 | 
    20  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    21  | 
  | 
| 
22092
 | 
    22  | 
	val scan_parm = OuterParse.name -- (OuterParse.$$$ "=" |-- OuterParse.name);
  | 
| 
14350
 | 
    23  | 
  | 
| 
22092
 | 
    24  | 
	val scan_parms = Scan.option
  | 
| 
26000
 | 
    25  | 
		(OuterParse.$$$ "[" |-- (OuterParse.list scan_parm) --| OuterParse.$$$ "]");
  | 
| 
14350
 | 
    26  | 
  | 
| 
 | 
    27  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    28  | 
(* set up the 'refute' command                                               *)
  | 
| 
 | 
    29  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    30  | 
  | 
| 
 | 
    31  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    32  | 
(* 'refute' takes an optional list of parameters, followed by an optional    *)
  | 
| 
 | 
    33  | 
(* subgoal number                                                            *)
  | 
| 
 | 
    34  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    35  | 
  | 
| 
 | 
    36  | 
	val refute_scan_args = scan_parms -- (Scan.option OuterParse.nat);
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    39  | 
(* the 'refute' command is mapped to 'Refute.refute_subgoal'                 *)
  | 
| 
 | 
    40  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
	fun refute_trans args =
  | 
| 
22092
 | 
    43  | 
		Toplevel.keep (fn state =>
  | 
| 
 | 
    44  | 
			let
  | 
| 
 | 
    45  | 
				val (parms_opt, subgoal_opt) = args
  | 
| 
 | 
    46  | 
				val parms   = Option.getOpt (parms_opt, [])
  | 
| 
 | 
    47  | 
				val subgoal = Option.getOpt (subgoal_opt, 1) - 1
  | 
| 
 | 
    48  | 
				val thy     = Toplevel.theory_of state
  | 
| 
 | 
    49  | 
				val thm     = (snd o snd) (Proof.get_goal (Toplevel.proof_of state))
  | 
| 
 | 
    50  | 
			in
  | 
| 
 | 
    51  | 
				Refute.refute_subgoal thy parms thm subgoal
  | 
| 
 | 
    52  | 
			end);
  | 
| 
14350
 | 
    53  | 
  | 
| 
 | 
    54  | 
	fun refute_parser tokens = (refute_scan_args tokens) |>> refute_trans;
  | 
| 
 | 
    55  | 
  | 
| 
24867
 | 
    56  | 
	val _ = OuterSyntax.improper_command "refute"
  | 
| 
22092
 | 
    57  | 
		"try to find a model that refutes a given subgoal"
  | 
| 
 | 
    58  | 
		OuterKeyword.diag refute_parser;
  | 
| 
14350
 | 
    59  | 
  | 
| 
 | 
    60  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    61  | 
(* set up the 'refute_params' command                                        *)
  | 
| 
 | 
    62  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    63  | 
  | 
| 
 | 
    64  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    65  | 
(* 'refute_params' takes an optional list of parameters                      *)
  | 
| 
 | 
    66  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    67  | 
  | 
| 
 | 
    68  | 
	val refute_params_scan_args = scan_parms;
  | 
| 
 | 
    69  | 
  | 
| 
 | 
    70  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    71  | 
(* the 'refute_params' command is mapped to (multiple calls of)              *)
  | 
| 
 | 
    72  | 
(* 'Refute.set_default_param'; then the (new) default parameters are         *)
  | 
| 
 | 
    73  | 
(* displayed                                                                 *)
  | 
| 
 | 
    74  | 
(* ------------------------------------------------------------------------- *)
  | 
| 
 | 
    75  | 
  | 
| 
 | 
    76  | 
	fun refute_params_trans args =
  | 
| 
 | 
    77  | 
		let
  | 
| 
22092
 | 
    78  | 
			fun add_params thy [] =
  | 
| 
 | 
    79  | 
				thy
  | 
| 
 | 
    80  | 
			  | add_params thy (p::ps) =
  | 
| 
 | 
    81  | 
				add_params (Refute.set_default_param p thy) ps
  | 
| 
14350
 | 
    82  | 
		in
  | 
| 
 | 
    83  | 
			Toplevel.theory (fn thy =>
  | 
| 
 | 
    84  | 
				let
  | 
| 
22092
 | 
    85  | 
					val thy'               = add_params thy (getOpt (args, []))
  | 
| 
14350
 | 
    86  | 
					val new_default_params = Refute.get_default_params thy'
  | 
| 
22092
 | 
    87  | 
					val output             = if new_default_params=[] then
  | 
| 
 | 
    88  | 
							"none"
  | 
| 
 | 
    89  | 
						else
  | 
| 
26931
 | 
    90  | 
							cat_lines (map (fn (name, value) => name ^ "=" ^ value)
  | 
| 
22092
 | 
    91  | 
								new_default_params)
  | 
| 
14350
 | 
    92  | 
				in
  | 
| 
 | 
    93  | 
					writeln ("Default parameters for 'refute':\n" ^ output);
 | 
| 
 | 
    94  | 
					thy'
  | 
| 
 | 
    95  | 
				end)
  | 
| 
 | 
    96  | 
		end;
  | 
| 
 | 
    97  | 
  | 
| 
22092
 | 
    98  | 
	fun refute_params_parser tokens =
  | 
| 
 | 
    99  | 
		(refute_params_scan_args tokens) |>> refute_params_trans;
  | 
| 
14350
 | 
   100  | 
  | 
| 
24867
 | 
   101  | 
	val _ = OuterSyntax.command "refute_params"
  | 
| 
22092
 | 
   102  | 
		"show/store default parameters for the 'refute' command"
  | 
| 
 | 
   103  | 
		OuterKeyword.thy_decl refute_params_parser;
  | 
| 
14350
 | 
   104  | 
  | 
| 
 | 
   105  | 
end;
  | 
| 
 | 
   106  | 
  |