| author | wenzelm | 
| Wed, 30 Mar 2011 23:26:40 +0200 | |
| changeset 42174 | d0be2722ce9f | 
| parent 42107 | a6725f293377 | 
| child 42270 | 5f2960582e45 | 
| permissions | -rw-r--r-- | 
| 39958 | 1 | (* Title: HOL/Tools/Metis/metis_reconstruct.ML | 
| 39495 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 2 | Author: Kong W. Susanto, Cambridge University Computer Laboratory | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 3 | Author: Lawrence C. Paulson, Cambridge University Computer Laboratory | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 4 | Author: Jasmin Blanchette, TU Muenchen | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 5 | Copyright Cambridge University 2007 | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 6 | |
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 7 | Proof reconstruction for Metis. | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 8 | *) | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 9 | |
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 10 | signature METIS_RECONSTRUCT = | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 11 | sig | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 12 | type mode = Metis_Translate.mode | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 13 | |
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 14 | val trace : bool Config.T | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 15 | val trace_msg : Proof.context -> (unit -> string) -> unit | 
| 40665 
1a65f0c74827
added "verbose" option to Metis to shut up its warnings if necessary
 blanchet parents: 
40264diff
changeset | 16 | val verbose : bool Config.T | 
| 
1a65f0c74827
added "verbose" option to Metis to shut up its warnings if necessary
 blanchet parents: 
40264diff
changeset | 17 | val verbose_warning : Proof.context -> string -> unit | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 18 | val lookth : (Metis_Thm.thm * 'a) list -> Metis_Thm.thm -> 'a | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 19 | val untyped_aconv : term -> term -> bool | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 20 | val replay_one_inference : | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 21 | Proof.context -> mode -> (string * term) list * int Unsynchronized.ref | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 22 | -> Metis_Thm.thm * Metis_Proof.inference -> (Metis_Thm.thm * thm) list | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 23 | -> (Metis_Thm.thm * thm) list | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 24 | val discharge_skolem_premises : | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 25 | Proof.context -> (thm * term) option list -> thm -> thm | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 26 | val setup : theory -> theory | 
| 39495 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 27 | end; | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 28 | |
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 29 | structure Metis_Reconstruct : METIS_RECONSTRUCT = | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 30 | struct | 
| 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 31 | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 32 | open Metis_Translate | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 33 | |
| 40724 
d01a1b3ab23d
renamed "trace_me{son,tis}" and "verbose_metis" to have the name of the tool first
 blanchet parents: 
40665diff
changeset | 34 | val (trace, trace_setup) = Attrib.config_bool "metis_trace" (K false) | 
| 
d01a1b3ab23d
renamed "trace_me{son,tis}" and "verbose_metis" to have the name of the tool first
 blanchet parents: 
40665diff
changeset | 35 | val (verbose, verbose_setup) = Attrib.config_bool "metis_verbose" (K true) | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 36 | |
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 37 | fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () | 
| 40665 
1a65f0c74827
added "verbose" option to Metis to shut up its warnings if necessary
 blanchet parents: 
40264diff
changeset | 38 | fun verbose_warning ctxt msg = | 
| 
1a65f0c74827
added "verbose" option to Metis to shut up its warnings if necessary
 blanchet parents: 
40264diff
changeset | 39 | if Config.get ctxt verbose then warning msg else () | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 40 | |
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 41 | datatype term_or_type = SomeTerm of term | SomeType of typ | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 42 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 43 | fun terms_of [] = [] | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 44 | | terms_of (SomeTerm t :: tts) = t :: terms_of tts | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 45 | | terms_of (SomeType _ :: tts) = terms_of tts; | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 46 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 47 | fun types_of [] = [] | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 48 | | types_of (SomeTerm (Var ((a,idx), _)) :: tts) = | 
| 42098 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 49 | if String.isPrefix metis_generated_var_prefix a then | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 50 | (*Variable generated by Metis, which might have been a type variable.*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 51 |           TVar (("'" ^ a, idx), HOLogic.typeS) :: types_of tts
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 52 | else types_of tts | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 53 | | types_of (SomeTerm _ :: tts) = types_of tts | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 54 | | types_of (SomeType T :: tts) = T :: types_of tts; | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 55 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 56 | fun apply_list rator nargs rands = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 57 | let val trands = terms_of rands | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 58 | in if length trands = nargs then SomeTerm (list_comb(rator, trands)) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 59 | else raise Fail | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 60 |         ("apply_list: wrong number of arguments: " ^ Syntax.string_of_term_global Pure.thy rator ^
 | 
| 41491 | 61 | " expected " ^ string_of_int nargs ^ | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 62 | " received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands)) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 63 | end; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 64 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 65 | fun infer_types ctxt = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 66 | Syntax.check_terms (ProofContext.set_mode ProofContext.mode_pattern ctxt); | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 67 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 68 | (*We use 1 rather than 0 because variable references in clauses may otherwise conflict | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 69 | with variable constraints in the goal...at least, type inference often fails otherwise. | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 70 | SEE ALSO axiom_inf below.*) | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 71 | fun mk_var (w, T) = Var ((w, 1), T) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 72 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 73 | (*include the default sort, if available*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 74 | fun mk_tfree ctxt w = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 75 | let val ww = "'" ^ w | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 76 | in TFree(ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1))) end; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 77 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 78 | (*Remove the "apply" operator from an HO term*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 79 | fun strip_happ args (Metis_Term.Fn(".",[t,u])) = strip_happ (u::args) t
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 80 | | strip_happ args x = (x, args); | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 81 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 82 | fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 83 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 84 | fun hol_type_from_metis_term _ (Metis_Term.Var v) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 85 | (case strip_prefix_and_unascii tvar_prefix v of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 86 | SOME w => make_tvar w | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 87 | | NONE => make_tvar v) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 88 | | hol_type_from_metis_term ctxt (Metis_Term.Fn(x, tys)) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 89 | (case strip_prefix_and_unascii type_const_prefix x of | 
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41139diff
changeset | 90 | SOME tc => Type (invert_const tc, | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 91 | map (hol_type_from_metis_term ctxt) tys) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 92 | | NONE => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 93 | case strip_prefix_and_unascii tfree_prefix x of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 94 | SOME tf => mk_tfree ctxt tf | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 95 |         | NONE    => raise Fail ("hol_type_from_metis_term: " ^ x));
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 96 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 97 | (*Maps metis terms to isabelle terms*) | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 98 | fun hol_term_from_metis_PT ctxt new_skolem_var_count fol_tm = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 99 | let val thy = ProofContext.theory_of ctxt | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 100 | val _ = trace_msg ctxt (fn () => "hol_term_from_metis_PT: " ^ | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 101 | Metis_Term.toString fol_tm) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 102 | fun tm_to_tt (Metis_Term.Var v) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 103 | (case strip_prefix_and_unascii tvar_prefix v of | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 104 | SOME w => SomeType (make_tvar w) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 105 | | NONE => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 106 | case strip_prefix_and_unascii schematic_var_prefix v of | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 107 | SOME w => SomeTerm (mk_var (w, HOLogic.typeT)) | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 108 | | NONE => SomeTerm (mk_var (v, HOLogic.typeT)) ) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 109 | (*Var from Metis with a name like _nnn; possibly a type variable*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 110 |         | tm_to_tt (Metis_Term.Fn ("{}", [arg])) = tm_to_tt arg   (*hBOOL*)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 111 |         | tm_to_tt (t as Metis_Term.Fn (".",_)) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 112 | let val (rator,rands) = strip_happ [] t | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 113 | in case rator of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 114 | Metis_Term.Fn(fname,ts) => applic_to_tt (fname, ts @ rands) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 115 | | _ => case tm_to_tt rator of | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 116 | SomeTerm t => SomeTerm (list_comb(t, terms_of (map tm_to_tt rands))) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 117 | | _ => raise Fail "tm_to_tt: HO application" | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 118 | end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 119 | | tm_to_tt (Metis_Term.Fn (fname, args)) = applic_to_tt (fname,args) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 120 |       and applic_to_tt ("=",ts) =
 | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 121 |             SomeTerm (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
 | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 122 | | applic_to_tt (a,ts) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 123 | case strip_prefix_and_unascii const_prefix a of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 124 | SOME b => | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 125 | let | 
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41139diff
changeset | 126 | val c = invert_const b | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 127 | val ntypes = num_type_args thy c | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 128 | val nterms = length ts - ntypes | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 129 | val tts = map tm_to_tt ts | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 130 | val tys = types_of (List.take(tts,ntypes)) | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 131 | val j = !new_skolem_var_count + 1 | 
| 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 132 | val _ = new_skolem_var_count := j | 
| 39939 
6e9aff5ee9b5
paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
 blanchet parents: 
39896diff
changeset | 133 | val t = | 
| 
6e9aff5ee9b5
paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
 blanchet parents: 
39896diff
changeset | 134 | if String.isPrefix new_skolem_const_prefix c then | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 135 | Var ((new_skolem_var_name_from_const c, j), | 
| 39939 
6e9aff5ee9b5
paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
 blanchet parents: 
39896diff
changeset | 136 | Type_Infer.paramify_vars (tl tys ---> hd tys)) | 
| 
6e9aff5ee9b5
paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
 blanchet parents: 
39896diff
changeset | 137 | else | 
| 
6e9aff5ee9b5
paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
 blanchet parents: 
39896diff
changeset | 138 | Const (c, dummyT) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 139 | in if length tys = ntypes then | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 140 | apply_list t nterms (List.drop(tts,ntypes)) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 141 | else | 
| 41491 | 142 |                        raise Fail ("Constant " ^ c ^ " expects " ^ string_of_int ntypes ^
 | 
| 143 | " but gets " ^ string_of_int (length tys) ^ | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 144 | " type arguments\n" ^ | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 145 | cat_lines (map (Syntax.string_of_typ ctxt) tys) ^ | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 146 | " the terms are \n" ^ | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 147 | cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts))) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 148 | end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 149 | | NONE => (*Not a constant. Is it a type constructor?*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 150 | case strip_prefix_and_unascii type_const_prefix a of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 151 | SOME b => | 
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41139diff
changeset | 152 | SomeType (Type (invert_const b, types_of (map tm_to_tt ts))) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 153 | | NONE => (*Maybe a TFree. Should then check that ts=[].*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 154 | case strip_prefix_and_unascii tfree_prefix a of | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 155 | SOME b => SomeType (mk_tfree ctxt b) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 156 | | NONE => (*a fixed variable? They are Skolem functions.*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 157 | case strip_prefix_and_unascii fixed_var_prefix a of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 158 | SOME b => | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 159 | let val opr = Free (b, HOLogic.typeT) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 160 | in apply_list opr (length ts) (map tm_to_tt ts) end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 161 |               | NONE => raise Fail ("unexpected metis function: " ^ a)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 162 | in | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 163 | case tm_to_tt fol_tm of | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 164 | SomeTerm t => t | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 165 |     | SomeType T => raise TYPE ("fol_tm_to_tt: Term expected", [T], [])
 | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 166 | end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 167 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 168 | (*Maps fully-typed metis terms to isabelle terms*) | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 169 | fun hol_term_from_metis_FT ctxt new_skolem_var_count fol_tm = | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 170 | let val _ = trace_msg ctxt (fn () => "hol_term_from_metis_FT: " ^ | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 171 | Metis_Term.toString fol_tm) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 172 |       fun cvt (Metis_Term.Fn ("ti", [Metis_Term.Var v, _])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 173 | (case strip_prefix_and_unascii schematic_var_prefix v of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 174 | SOME w => mk_var(w, dummyT) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 175 | | NONE => mk_var(v, dummyT)) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 176 |         | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn ("=",[]), _])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 177 |             Const (@{const_name HOL.eq}, HOLogic.typeT)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 178 |         | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (x,[]), ty])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 179 | (case strip_prefix_and_unascii const_prefix x of | 
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41139diff
changeset | 180 | SOME c => Const (invert_const c, dummyT) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 181 | | NONE => (*Not a constant. Is it a fixed variable??*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 182 | case strip_prefix_and_unascii fixed_var_prefix x of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 183 | SOME v => Free (v, hol_type_from_metis_term ctxt ty) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 184 |               | NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 185 |         | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (".",[tm1,tm2]), _])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 186 | cvt tm1 $ cvt tm2 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 187 |         | cvt (Metis_Term.Fn (".",[tm1,tm2])) = (*untyped application*)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 188 | cvt tm1 $ cvt tm2 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 189 |         | cvt (Metis_Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 190 |         | cvt (Metis_Term.Fn ("=", [tm1,tm2])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 191 |             list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 192 | | cvt (t as Metis_Term.Fn (x, [])) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 193 | (case strip_prefix_and_unascii const_prefix x of | 
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41139diff
changeset | 194 | SOME c => Const (invert_const c, dummyT) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 195 | | NONE => (*Not a constant. Is it a fixed variable??*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 196 | case strip_prefix_and_unascii fixed_var_prefix x of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 197 | SOME v => Free (v, dummyT) | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 198 | | NONE => (trace_msg ctxt (fn () => | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 199 | "hol_term_from_metis_FT bad const: " ^ x); | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 200 | hol_term_from_metis_PT ctxt new_skolem_var_count t)) | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 201 | | cvt t = (trace_msg ctxt (fn () => | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 202 | "hol_term_from_metis_FT bad term: " ^ Metis_Term.toString t); | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 203 | hol_term_from_metis_PT ctxt new_skolem_var_count t) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 204 | in fol_tm |> cvt end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 205 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 206 | fun hol_term_from_metis FT = hol_term_from_metis_FT | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 207 | | hol_term_from_metis _ = hol_term_from_metis_PT | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 208 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 209 | fun hol_terms_from_metis ctxt mode (old_skolems, new_skolem_var_count) fol_tms = | 
| 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 210 | let val ts = map (hol_term_from_metis mode ctxt new_skolem_var_count) fol_tms | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 211 | val _ = trace_msg ctxt (fn () => " calling type inference:") | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 212 | val _ = app (fn t => trace_msg ctxt | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 213 | (fn () => Syntax.string_of_term ctxt t)) ts | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 214 | val ts' = ts |> map (reveal_old_skolem_terms old_skolems) | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 215 | |> infer_types ctxt | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 216 | val _ = app (fn t => trace_msg ctxt | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 217 | (fn () => " final term: " ^ Syntax.string_of_term ctxt t ^ | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 218 | " of type " ^ Syntax.string_of_typ ctxt (type_of t))) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 219 | ts' | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 220 | in ts' end; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 221 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 222 | (* ------------------------------------------------------------------------- *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 223 | (* FOL step Inference Rules *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 224 | (* ------------------------------------------------------------------------- *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 225 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 226 | (*for debugging only*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 227 | (* | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 228 | fun print_thpair ctxt (fth,th) = | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 229 | (trace_msg ctxt (fn () => "============================================="); | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 230 | trace_msg ctxt (fn () => "Metis: " ^ Metis_Thm.toString fth); | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 231 | trace_msg ctxt (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th)); | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 232 | *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 233 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 234 | fun lookth thpairs (fth : Metis_Thm.thm) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 235 | the (AList.lookup (uncurry Metis_Thm.equal) thpairs fth) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 236 | handle Option.Option => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 237 |          raise Fail ("Failed to find Metis theorem " ^ Metis_Thm.toString fth)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 238 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 239 | fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx)); | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 240 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 241 | (* INFERENCE RULE: AXIOM *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 242 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 243 | fun axiom_inf thpairs th = Thm.incr_indexes 1 (lookth thpairs th); | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 244 | (*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 245 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 246 | (* INFERENCE RULE: ASSUME *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 247 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 248 | val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 249 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 250 | fun inst_excluded_middle thy i_atm = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 251 | let val th = EXCLUDED_MIDDLE | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 252 | val [vx] = Term.add_vars (prop_of th) [] | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 253 | val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)] | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 254 | in cterm_instantiate substs th end; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 255 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 256 | fun assume_inf ctxt mode skolem_params atm = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 257 | inst_excluded_middle | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 258 | (ProofContext.theory_of ctxt) | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 259 | (singleton (hol_terms_from_metis ctxt mode skolem_params) | 
| 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 260 | (Metis_Term.Fn atm)) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 261 | |
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 262 | (* INFERENCE RULE: INSTANTIATE (SUBST). Type instantiations are ignored. Trying | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 263 | to reconstruct them admits new possibilities of errors, e.g. concerning | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 264 | sorts. Instead we try to arrange that new TVars are distinct and that types | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 265 | can be inferred from terms. *) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 266 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 267 | fun inst_inf ctxt mode (old_skolems, new_skolem_var_count) thpairs fsubst th = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 268 | let val thy = ProofContext.theory_of ctxt | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 269 | val i_th = lookth thpairs th | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 270 | val i_th_vars = Term.add_vars (prop_of i_th) [] | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 271 | fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 272 | fun subst_translation (x,y) = | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 273 | let val v = find_var x | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 274 | (* We call "reveal_old_skolem_terms" and "infer_types" below. *) | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 275 | val t = hol_term_from_metis mode ctxt new_skolem_var_count y | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 276 | in SOME (cterm_of thy (Var v), t) end | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 277 | handle Option.Option => | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 278 | (trace_msg ctxt (fn () => "\"find_var\" failed for " ^ x ^ | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 279 | " in " ^ Display.string_of_thm ctxt i_th); | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 280 | NONE) | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 281 | | TYPE _ => | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 282 | (trace_msg ctxt (fn () => "\"hol_term_from_metis\" failed for " ^ x ^ | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 283 | " in " ^ Display.string_of_thm ctxt i_th); | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 284 | NONE) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 285 | fun remove_typeinst (a, t) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 286 | case strip_prefix_and_unascii schematic_var_prefix a of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 287 | SOME b => SOME (b, t) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 288 | | NONE => case strip_prefix_and_unascii tvar_prefix a of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 289 | SOME _ => NONE (*type instantiations are forbidden!*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 290 | | NONE => SOME (a,t) (*internal Metis var?*) | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 291 | val _ = trace_msg ctxt (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 292 | val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 293 | val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs) | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 294 | val tms = rawtms |> map (reveal_old_skolem_terms old_skolems) | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39550diff
changeset | 295 | |> infer_types ctxt | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 296 | val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 297 | val substs' = ListPair.zip (vars, map ctm_of tms) | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 298 | val _ = trace_msg ctxt (fn () => | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 299 |         cat_lines ("subst_translations:" ::
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 300 | (substs' |> map (fn (x, y) => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 301 | Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^ | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 302 | Syntax.string_of_term ctxt (term_of y))))); | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 303 | in cterm_instantiate substs' i_th end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 304 | handle THM (msg, _, _) => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 305 |          error ("Cannot replay Metis proof in Isabelle:\n" ^ msg)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 306 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 307 | (* INFERENCE RULE: RESOLVE *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 308 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 309 | (* Like RSN, but we rename apart only the type variables. Vars here typically | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 310 | have an index of 1, and the use of RSN would increase this typically to 3. | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 311 | Instantiations of those Vars could then fail. See comment on "mk_var". *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 312 | fun resolve_inc_tyvars thy tha i thb = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 313 | let | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 314 | val tha = Drule.incr_type_indexes (1 + Thm.maxidx_of thb) tha | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 315 | fun aux tha thb = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 316 | case Thm.bicompose false (false, tha, nprems_of tha) i thb | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 317 | |> Seq.list_of |> distinct Thm.eq_thm of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 318 | [th] => th | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 319 |       | _ => raise THM ("resolve_inc_tyvars: unique result expected", i,
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 320 | [tha, thb]) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 321 | in | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 322 | aux tha thb | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 323 | handle TERM z => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 324 | (* The unifier, which is invoked from "Thm.bicompose", will sometimes | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 325 | refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 326 | "TERM" exception (with "add_ffpair" as first argument). We then | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 327 | perform unification of the types of variables by hand and try | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 328 | again. We could do this the first time around but this error | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 329 | occurs seldom and we don't want to break existing proofs in subtle | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 330 | ways or slow them down needlessly. *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 331 | case [] |> fold (Term.add_vars o prop_of) [tha, thb] | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 332 | |> AList.group (op =) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 333 | |> maps (fn ((s, _), T :: Ts) => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 334 | map (fn T' => (Free (s, T), Free (s, T'))) Ts) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 335 | |> rpair (Envir.empty ~1) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 336 | |-> fold (Pattern.unify thy) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 337 | |> Envir.type_env |> Vartab.dest | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 338 | |> map (fn (x, (S, T)) => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 339 | pairself (ctyp_of thy) (TVar (x, S), T)) of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 340 | [] => raise TERM z | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 341 | | ps => aux (instantiate (ps, []) tha) (instantiate (ps, []) thb) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 342 | end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 343 | |
| 40221 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 344 | fun s_not (@{const Not} $ t) = t
 | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 345 | | s_not t = HOLogic.mk_not t | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 346 | fun simp_not_not (@{const Not} $ t) = s_not (simp_not_not t)
 | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 347 | | simp_not_not t = t | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 348 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 349 | (* Match untyped terms. *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 350 | fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 351 | | untyped_aconv (Free (a, _)) (Free (b, _)) = (a = b) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 352 | | untyped_aconv (Var ((a, _), _)) (Var ((b, _), _)) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 353 | (a = b) (* The index is ignored, for some reason. *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 354 | | untyped_aconv (Bound i) (Bound j) = (i = j) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 355 | | untyped_aconv (Abs (_, _, t)) (Abs (_, _, u)) = untyped_aconv t u | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 356 | | untyped_aconv (t1 $ t2) (u1 $ u2) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 357 | untyped_aconv t1 u1 andalso untyped_aconv t2 u2 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 358 | | untyped_aconv _ _ = false | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 359 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 360 | (* Finding the relative location of an untyped term within a list of terms *) | 
| 40221 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 361 | fun index_of_literal lit haystack = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 362 | let | 
| 40221 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 363 | val normalize = simp_not_not o Envir.eta_contract | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 364 | val match_lit = | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 365 | HOLogic.dest_Trueprop #> normalize #> untyped_aconv (lit |> normalize) | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 366 | in case find_index match_lit haystack of ~1 => raise Empty | n => n + 1 end | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 367 | |
| 39893 | 368 | (* Permute a rule's premises to move the i-th premise to the last position. *) | 
| 369 | fun make_last i th = | |
| 370 | let val n = nprems_of th | |
| 371 | in if 1 <= i andalso i <= n | |
| 372 | then Thm.permute_prems (i-1) 1 th | |
| 373 |       else raise THM("select_literal", i, [th])
 | |
| 374 | end; | |
| 375 | ||
| 41143 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 376 | val neg_neg_imp = @{lemma "~ ~ Q ==> Q" by blast}
 | 
| 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 377 | |
| 39893 | 378 | (* Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing | 
| 379 | double-negations. *) | |
| 41143 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 380 | fun negate_head thy = | 
| 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 381 |   rewrite_rule @{thms atomize_not}
 | 
| 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 382 | #> perhaps (try (fn th => resolve_inc_tyvars thy th 1 neg_neg_imp)) | 
| 39893 | 383 | |
| 384 | (* Maps the clause [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P *) | |
| 41143 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 385 | fun select_literal thy = negate_head thy oo make_last | 
| 39893 | 386 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 387 | fun resolve_inf ctxt mode skolem_params thpairs atm th1 th2 = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 388 | let | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 389 | val thy = ProofContext.theory_of ctxt | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 390 | val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2 | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 391 | val _ = trace_msg ctxt (fn () => " isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1) | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 392 | val _ = trace_msg ctxt (fn () => " isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 393 | in | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 394 | (* Trivial cases where one operand is type info *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 395 | if Thm.eq_thm (TrueI, i_th1) then | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 396 | i_th2 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 397 | else if Thm.eq_thm (TrueI, i_th2) then | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 398 | i_th1 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 399 | else | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 400 | let | 
| 40221 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 401 | val i_atm = | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 402 | singleton (hol_terms_from_metis ctxt mode skolem_params) | 
| 40221 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 403 | (Metis_Term.Fn atm) | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 404 | val _ = trace_msg ctxt (fn () => " atom: " ^ Syntax.string_of_term ctxt i_atm) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 405 | val prems_th1 = prems_of i_th1 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 406 | val prems_th2 = prems_of i_th2 | 
| 40221 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 407 | val index_th1 = | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 408 | index_of_literal (s_not i_atm) prems_th1 | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 409 | handle Empty => raise Fail "Failed to find literal in th1" | 
| 41491 | 410 | val _ = trace_msg ctxt (fn () => " index_th1: " ^ string_of_int index_th1) | 
| 40221 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 411 | val index_th2 = | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 412 | index_of_literal i_atm prems_th2 | 
| 
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
 blanchet parents: 
40158diff
changeset | 413 | handle Empty => raise Fail "Failed to find literal in th2" | 
| 41491 | 414 | val _ = trace_msg ctxt (fn () => " index_th2: " ^ string_of_int index_th2) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 415 | in | 
| 41143 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 416 | resolve_inc_tyvars thy (select_literal thy index_th1 i_th1) index_th2 | 
| 
0b05cc14c2cb
remove at most one double negation -- any other double negations are part of some higher-order reasoning and should be left alone, cf. "HO_Reas.thy"
 blanchet parents: 
41140diff
changeset | 417 | i_th2 | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 418 | end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 419 | end; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 420 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 421 | (* INFERENCE RULE: REFL *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 422 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 423 | val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 424 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 425 | val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 426 | val refl_idx = 1 + Thm.maxidx_of REFL_THM; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 427 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 428 | fun refl_inf ctxt mode skolem_params t = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 429 | let val thy = ProofContext.theory_of ctxt | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 430 | val i_t = singleton (hol_terms_from_metis ctxt mode skolem_params) t | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 431 | val _ = trace_msg ctxt (fn () => " term: " ^ Syntax.string_of_term ctxt i_t) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 432 | val c_t = cterm_incr_types thy refl_idx i_t | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 433 | in cterm_instantiate [(refl_x, c_t)] REFL_THM end; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 434 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 435 | (* INFERENCE RULE: EQUALITY *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 436 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 437 | val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 438 | val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 439 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 440 | val metis_eq = Metis_Term.Fn ("=", []);
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 441 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 442 | fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0  (*equality has no type arguments*)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 443 | | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 444 | | get_ty_arg_size _ _ = 0; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 445 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 446 | fun equality_inf ctxt mode skolem_params (pos, atm) fp fr = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 447 | let val thy = ProofContext.theory_of ctxt | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 448 | val m_tm = Metis_Term.Fn atm | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 449 | val [i_atm,i_tm] = hol_terms_from_metis ctxt mode skolem_params [m_tm, fr] | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 450 | val _ = trace_msg ctxt (fn () => "sign of the literal: " ^ Bool.toString pos) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 451 | fun replace_item_list lx 0 (_::ls) = lx::ls | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 452 | | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 453 | fun path_finder_FO tm [] = (tm, Bound 0) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 454 | | path_finder_FO tm (p::ps) = | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 455 | let val (tm1,args) = strip_comb tm | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 456 | val adjustment = get_ty_arg_size thy tm1 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 457 | val p' = if adjustment > p then p else p-adjustment | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 458 | val tm_p = List.nth(args,p') | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 459 | handle Subscript => | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 460 |                          error ("Cannot replay Metis proof in Isabelle:\n" ^
 | 
| 41491 | 461 | "equality_inf: " ^ string_of_int p ^ " adj " ^ | 
| 462 | string_of_int adjustment ^ " term " ^ | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 463 | Syntax.string_of_term ctxt tm) | 
| 41491 | 464 | val _ = trace_msg ctxt (fn () => "path_finder: " ^ string_of_int p ^ | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 465 | " " ^ Syntax.string_of_term ctxt tm_p) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 466 | val (r,t) = path_finder_FO tm_p ps | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 467 | in | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 468 | (r, list_comb (tm1, replace_item_list t p' args)) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 469 | end | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 470 | fun path_finder_HO tm [] = (tm, Bound 0) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 471 | | path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 472 | | path_finder_HO (t$u) (_::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 473 | | path_finder_HO tm ps = | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 474 |           raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
 | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 475 | "equality_inf, path_finder_HO: path = " ^ | 
| 41491 | 476 | space_implode " " (map string_of_int ps) ^ | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 477 | " isa-term: " ^ Syntax.string_of_term ctxt tm) | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 478 | fun path_finder_FT tm [] _ = (tm, Bound 0) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 479 |         | path_finder_FT tm (0::ps) (Metis_Term.Fn ("ti", [t1, _])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 480 | path_finder_FT tm ps t1 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 481 |         | path_finder_FT (t$u) (0::ps) (Metis_Term.Fn (".", [t1, _])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 482 | (fn(x,y) => (x, y$u)) (path_finder_FT t ps t1) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 483 |         | path_finder_FT (t$u) (1::ps) (Metis_Term.Fn (".", [_, t2])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 484 | (fn(x,y) => (x, t$y)) (path_finder_FT u ps t2) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 485 | | path_finder_FT tm ps t = | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 486 |           raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
 | 
| 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 487 | "equality_inf, path_finder_FT: path = " ^ | 
| 41491 | 488 | space_implode " " (map string_of_int ps) ^ | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 489 | " isa-term: " ^ Syntax.string_of_term ctxt tm ^ | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 490 | " fol-term: " ^ Metis_Term.toString t) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 491 | fun path_finder FO tm ps _ = path_finder_FO tm ps | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 492 |         | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 493 | (*equality: not curried, as other predicates are*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 494 | if p=0 then path_finder_HO tm (0::1::ps) (*select first operand*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 495 | else path_finder_HO tm (p::ps) (*1 selects second operand*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 496 |         | path_finder HO tm (_ :: ps) (Metis_Term.Fn ("{}", [_])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 497 | path_finder_HO tm ps (*if not equality, ignore head to skip hBOOL*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 498 |         | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 499 |                             (Metis_Term.Fn ("=", [t1,t2])) =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 500 | (*equality: not curried, as other predicates are*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 501 | if p=0 then path_finder_FT tm (0::1::ps) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 502 |                           (Metis_Term.Fn (".", [Metis_Term.Fn (".", [metis_eq,t1]), t2]))
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 503 | (*select first operand*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 504 | else path_finder_FT tm (p::ps) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 505 |                    (Metis_Term.Fn (".", [metis_eq,t2]))
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 506 | (*1 selects second operand*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 507 |         | path_finder FT tm (_ :: ps) (Metis_Term.Fn ("{}", [t1])) = path_finder_FT tm ps t1
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 508 | (*if not equality, ignore head to skip the hBOOL predicate*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 509 | | path_finder FT tm ps t = path_finder_FT tm ps t (*really an error case!*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 510 |       fun path_finder_lit ((nt as Const (@{const_name Not}, _)) $ tm_a) idx =
 | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 511 | let val (tm, tm_rslt) = path_finder mode tm_a idx m_tm | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 512 | in (tm, nt $ tm_rslt) end | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 513 | | path_finder_lit tm_a idx = path_finder mode tm_a idx m_tm | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 514 | val (tm_subst, body) = path_finder_lit i_atm fp | 
| 39498 
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
 blanchet parents: 
39497diff
changeset | 515 |       val tm_abs = Abs ("x", type_of tm_subst, body)
 | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 516 | val _ = trace_msg ctxt (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs) | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 517 | val _ = trace_msg ctxt (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm) | 
| 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 518 | val _ = trace_msg ctxt (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 519 | val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill typed but gives right max*) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 520 | val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em) | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 521 | val _ = trace_msg ctxt (fn () => "subst' " ^ Display.string_of_thm ctxt subst') | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 522 | val eq_terms = map (pairself (cterm_of thy)) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 523 | (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm])) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 524 | in cterm_instantiate eq_terms subst' end; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 525 | |
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 526 | val factor = Seq.hd o distinct_subgoals_tac; | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 527 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 528 | fun step ctxt mode skolem_params thpairs p = | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 529 | case p of | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 530 | (fol_th, Metis_Proof.Axiom _) => factor (axiom_inf thpairs fol_th) | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 531 | | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode skolem_params f_atm | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 532 | | (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) => | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 533 | factor (inst_inf ctxt mode skolem_params thpairs f_subst f_th1) | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 534 | | (_, Metis_Proof.Resolve(f_atm, f_th1, f_th2)) => | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 535 | factor (resolve_inf ctxt mode skolem_params thpairs f_atm f_th1 f_th2) | 
| 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 536 | | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode skolem_params f_tm | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 537 | | (_, Metis_Proof.Equality (f_lit, f_p, f_r)) => | 
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 538 | equality_inf ctxt mode skolem_params f_lit f_p f_r | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 539 | |
| 39893 | 540 | fun flexflex_first_order th = | 
| 541 | case Thm.tpairs_of th of | |
| 542 | [] => th | |
| 543 | | pairs => | |
| 544 | let val thy = theory_of_thm th | |
| 545 | val (_, tenv) = | |
| 546 | fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty) | |
| 547 | val t_pairs = map Meson.term_pair_of (Vartab.dest tenv) | |
| 548 | val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th | |
| 549 | in th' end | |
| 550 | handle THM _ => th; | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 551 | |
| 39895 
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
 blanchet parents: 
39893diff
changeset | 552 | fun is_metis_literal_genuine (_, (s, _)) = not (String.isPrefix class_prefix s) | 
| 
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
 blanchet parents: 
39893diff
changeset | 553 | fun is_isabelle_literal_genuine t = | 
| 39953 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 blanchet parents: 
39946diff
changeset | 554 |   case t of _ $ (Const (@{const_name Meson.skolem}, _) $ _) => false | _ => true
 | 
| 39895 
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
 blanchet parents: 
39893diff
changeset | 555 | |
| 
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
 blanchet parents: 
39893diff
changeset | 556 | fun count p xs = fold (fn x => if p x then Integer.add 1 else I) xs 0 | 
| 
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
 blanchet parents: 
39893diff
changeset | 557 | |
| 40259 
c0e34371c2e2
prevent type errors because of inconsistent skolem Var types by giving fresh indices to Skolems
 blanchet parents: 
40258diff
changeset | 558 | fun replay_one_inference ctxt mode skolem_params (fol_th, inf) thpairs = | 
| 40868 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 559 |   if not (null thpairs) andalso prop_of (snd (hd thpairs)) aconv @{prop False} then
 | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 560 | (* Isabelle sometimes identifies literals (premises) that are distinct in | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 561 | Metis (e.g., because of type variables). We give the Isabelle proof the | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 562 | benefice of the doubt. *) | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 563 | thpairs | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 564 | else | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 565 | let | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 566 | val _ = trace_msg ctxt | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 567 | (fn () => "=============================================") | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 568 | val _ = trace_msg ctxt | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 569 | (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th) | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 570 | val _ = trace_msg ctxt | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 571 | (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf) | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 572 | val th = step ctxt mode skolem_params thpairs (fol_th, inf) | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 573 | |> flexflex_first_order | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 574 | val _ = trace_msg ctxt | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 575 | (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th) | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 576 | val _ = trace_msg ctxt | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 577 | (fn () => "=============================================") | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 578 | val num_metis_lits = | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 579 | count is_metis_literal_genuine | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 580 | (Metis_LiteralSet.toList (Metis_Thm.clause fol_th)) | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 581 | val num_isabelle_lits = count is_isabelle_literal_genuine (prems_of th) | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 582 | val _ = if num_metis_lits >= num_isabelle_lits then () | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 583 | else error "Cannot replay Metis proof in Isabelle: Out of sync." | 
| 
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
 blanchet parents: 
40724diff
changeset | 584 | in (fol_th, th) :: thpairs end | 
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39495diff
changeset | 585 | |
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 586 | fun term_instantiate thy = cterm_instantiate o map (pairself (cterm_of thy)) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 587 | |
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 588 | val copy_prem = @{lemma "P ==> (P ==> P ==> Q) ==> Q" by fast}
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 589 | |
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 590 | fun copy_prems_tac [] ns i = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 591 | if forall (curry (op =) 1) ns then all_tac else copy_prems_tac (rev ns) [] i | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 592 | | copy_prems_tac (1 :: ms) ns i = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 593 | rotate_tac 1 i THEN copy_prems_tac ms (1 :: ns) i | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 594 | | copy_prems_tac (m :: ms) ns i = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 595 | etac copy_prem i THEN copy_prems_tac ms (m div 2 :: (m + 1) div 2 :: ns) i | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 596 | |
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 597 | fun instantiate_forall_tac thy t i st = | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 598 | let | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 599 | val params = Logic.strip_params (Logic.get_goal (prop_of st) i) |> rev | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 600 | fun repair (t as (Var ((s, _), _))) = | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 601 | (case find_index (fn (s', _) => s' = s) params of | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 602 | ~1 => t | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 603 | | j => Bound j) | 
| 40261 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 604 | | repair (t $ u) = | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 605 | (case (repair t, repair u) of | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 606 | (t as Bound j, u as Bound k) => | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 607 | (* This is a rather subtle trick to repair the discrepancy between | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 608 | the fully skolemized term that MESON gives us (where existentials | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 609 | were pulled out) and the reality. *) | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 610 | if k > j then t else t $ u | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 611 | | (t, u) => t $ u) | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 612 | | repair t = t | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 613 | val t' = t |> repair |> fold (curry absdummy) (map snd params) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 614 | fun do_instantiate th = | 
| 42098 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 615 | let | 
| 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 616 | val var = Term.add_vars (prop_of th) [] | 
| 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 617 | |> the_single | 
| 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 618 | in th |> term_instantiate thy [(Var var, t')] end | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 619 | in | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 620 |     (etac @{thm allE} i
 | 
| 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 621 | THEN rotate_tac ~1 i | 
| 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 622 | THEN PRIMITIVE do_instantiate) st | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 623 | end | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 624 | |
| 41135 | 625 | fun fix_exists_tac t = | 
| 40261 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 626 |   etac @{thm exE}
 | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 627 | THEN' rename_tac [t |> dest_Var |> fst |> fst] | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 628 | |
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 629 | fun release_quantifier_tac thy (skolem, t) = | 
| 41135 | 630 | (if skolem then fix_exists_tac else instantiate_forall_tac thy) t | 
| 40261 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 631 | |
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 632 | fun release_clusters_tac _ _ _ [] = K all_tac | 
| 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 633 | | release_clusters_tac thy ax_counts substs | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 634 | ((ax_no, cluster_no) :: clusters) = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 635 | let | 
| 40261 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 636 | val cluster_of_var = | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 637 | Meson_Clausify.cluster_of_zapped_var_name o fst o fst o dest_Var | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 638 | fun in_right_cluster ((_, (cluster_no', _)), _) = cluster_no' = cluster_no | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 639 | val cluster_substs = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 640 | substs | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 641 | |> map_filter (fn (ax_no', (_, (_, tsubst))) => | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 642 | if ax_no' = ax_no then | 
| 40261 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 643 | tsubst |> map (apfst cluster_of_var) | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 644 | |> filter (in_right_cluster o fst) | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 645 | |> map (apfst snd) | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 646 | |> SOME | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 647 | else | 
| 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 648 | NONE) | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 649 | fun do_cluster_subst cluster_subst = | 
| 40261 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40259diff
changeset | 650 | map (release_quantifier_tac thy) cluster_subst @ [rotate_tac 1] | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 651 | val first_prem = find_index (fn (ax_no', _) => ax_no' = ax_no) substs | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 652 | in | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 653 | rotate_tac first_prem | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 654 | THEN' (EVERY' (maps do_cluster_subst cluster_substs)) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 655 | THEN' rotate_tac (~ first_prem - length cluster_substs) | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 656 | THEN' release_clusters_tac thy ax_counts substs clusters | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 657 | end | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 658 | |
| 40264 
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
 blanchet parents: 
40261diff
changeset | 659 | fun cluster_key ((ax_no, (cluster_no, index_no)), skolem) = | 
| 
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
 blanchet parents: 
40261diff
changeset | 660 | (ax_no, (cluster_no, (skolem, index_no))) | 
| 
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
 blanchet parents: 
40261diff
changeset | 661 | fun cluster_ord p = | 
| 
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
 blanchet parents: 
40261diff
changeset | 662 | prod_ord int_ord (prod_ord int_ord (prod_ord bool_ord int_ord)) | 
| 
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
 blanchet parents: 
40261diff
changeset | 663 | (pairself cluster_key p) | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 664 | |
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 665 | val tysubst_ord = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 666 | list_ord (prod_ord Term_Ord.fast_indexname_ord | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 667 | (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord)) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 668 | |
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 669 | structure Int_Tysubst_Table = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 670 | Table(type key = int * (indexname * (sort * typ)) list | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 671 | val ord = prod_ord int_ord tysubst_ord) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 672 | |
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 673 | structure Int_Pair_Graph = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 674 | Graph(type key = int * int val ord = prod_ord int_ord int_ord) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 675 | |
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 676 | fun shuffle_key (((axiom_no, (_, index_no)), _), _) = (index_no, axiom_no) | 
| 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 677 | fun shuffle_ord p = prod_ord int_ord int_ord (pairself shuffle_key p) | 
| 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 678 | |
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 679 | (* Attempts to derive the theorem "False" from a theorem of the form | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 680 | "P1 ==> ... ==> Pn ==> False", where the "Pi"s are to be discharged using the | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 681 | specified axioms. The axioms have leading "All" and "Ex" quantifiers, which | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 682 | must be eliminated first. *) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 683 | fun discharge_skolem_premises ctxt axioms prems_imp_false = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 684 |   if prop_of prems_imp_false aconv @{prop False} then
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 685 | prems_imp_false | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 686 | else | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 687 | let | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 688 | val thy = ProofContext.theory_of ctxt | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 689 | fun match_term p = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 690 | let | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 691 | val (tyenv, tenv) = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 692 | Pattern.first_order_match thy p (Vartab.empty, Vartab.empty) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 693 | val tsubst = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 694 | tenv |> Vartab.dest | 
| 42099 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 695 | |> filter (Meson_Clausify.is_zapped_var_name o fst o fst) | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 696 | |> sort (cluster_ord | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 697 | o pairself (Meson_Clausify.cluster_of_zapped_var_name | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 698 | o fst o fst)) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 699 | |> map (Meson.term_pair_of | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 700 | #> pairself (Envir.subst_term_types tyenv)) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 701 | val tysubst = tyenv |> Vartab.dest | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 702 | in (tysubst, tsubst) end | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 703 | fun subst_info_for_prem subgoal_no prem = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 704 | case prem of | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 705 |           _ $ (Const (@{const_name Meson.skolem}, _) $ (_ $ t $ num)) =>
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 706 | let val ax_no = HOLogic.dest_nat num in | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 707 | (ax_no, (subgoal_no, | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 708 | match_term (nth axioms ax_no |> the |> snd, t))) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 709 | end | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 710 |         | _ => raise TERM ("discharge_skolem_premises: Malformed premise",
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 711 | [prem]) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 712 | fun cluster_of_var_name skolem s = | 
| 42098 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 713 | case try Meson_Clausify.cluster_of_zapped_var_name s of | 
| 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 714 | NONE => NONE | 
| 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 715 | | SOME ((ax_no, (cluster_no, _)), skolem') => | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 716 | if skolem' = skolem andalso cluster_no > 0 then | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 717 | SOME (ax_no, cluster_no) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 718 | else | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 719 | NONE | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 720 | fun clusters_in_term skolem t = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 721 | Term.add_var_names t [] |> map_filter (cluster_of_var_name skolem o fst) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 722 | fun deps_for_term_subst (var, t) = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 723 | case clusters_in_term false var of | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 724 | [] => NONE | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 725 | | [(ax_no, cluster_no)] => | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 726 | SOME ((ax_no, cluster_no), | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 727 | clusters_in_term true t | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 728 | |> cluster_no > 1 ? cons (ax_no, cluster_no - 1)) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 729 |         | _ => raise TERM ("discharge_skolem_premises: Expected Var", [var])
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 730 | val prems = Logic.strip_imp_prems (prop_of prems_imp_false) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 731 | val substs = prems |> map2 subst_info_for_prem (1 upto length prems) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 732 | |> sort (int_ord o pairself fst) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 733 | val depss = maps (map_filter deps_for_term_subst o snd o snd o snd) substs | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 734 | val clusters = maps (op ::) depss | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 735 | val ordered_clusters = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 736 | Int_Pair_Graph.empty | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 737 | |> fold Int_Pair_Graph.default_node (map (rpair ()) clusters) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 738 | |> fold Int_Pair_Graph.add_deps_acyclic depss | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 739 | |> Int_Pair_Graph.topological_order | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 740 | handle Int_Pair_Graph.CYCLES _ => | 
| 40158 | 741 | error "Cannot replay Metis proof in Isabelle without \ | 
| 742 | \\"Hilbert_Choice\"." | |
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 743 | val outer_param_names = | 
| 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 744 | [] |> fold (Term.add_var_names o snd) (map_filter I axioms) | 
| 42099 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 745 | |> filter (Meson_Clausify.is_zapped_var_name o fst) | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 746 | |> map (`(Meson_Clausify.cluster_of_zapped_var_name o fst)) | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 747 | |> filter (fn (((_, (cluster_no, _)), skolem), _) => | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 748 | cluster_no = 0 andalso skolem) | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 749 | |> sort shuffle_ord |> map (fst o snd) | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 750 | val ax_counts = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 751 | Int_Tysubst_Table.empty | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 752 | |> fold (fn (ax_no, (_, (tysubst, _))) => | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 753 | Int_Tysubst_Table.map_default ((ax_no, tysubst), 0) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 754 | (Integer.add 1)) substs | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 755 | |> Int_Tysubst_Table.dest | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 756 | (* for debugging: | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 757 | fun string_for_subst_info (ax_no, (subgoal_no, (tysubst, tsubst))) = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 758 | "ax: " ^ string_of_int ax_no ^ "; asm: " ^ string_of_int subgoal_no ^ | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 759 |         "; tysubst: " ^ PolyML.makestring tysubst ^ "; tsubst: {" ^
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 760 | commas (map ((fn (s, t) => s ^ " |-> " ^ t) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 761 | o pairself (Syntax.string_of_term ctxt)) tsubst) ^ "}" | 
| 40264 
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
 blanchet parents: 
40261diff
changeset | 762 |       val _ = tracing ("ORDERED CLUSTERS: " ^ PolyML.makestring ordered_clusters)
 | 
| 
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
 blanchet parents: 
40261diff
changeset | 763 |       val _ = tracing ("AXIOM COUNTS: " ^ PolyML.makestring ax_counts)
 | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 764 |       val _ = tracing ("SUBSTS (" ^ string_of_int (length substs) ^ "):\n" ^
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 765 | cat_lines (map string_for_subst_info substs)) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 766 | *) | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 767 | fun rotation_for_subgoal i = | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 768 | find_index (fn (_, (subgoal_no, _)) => subgoal_no = i) substs | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 769 | in | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 770 |       Goal.prove ctxt [] [] @{prop False}
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 771 | (K (cut_rules_tac | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 772 | (map (fst o the o nth axioms o fst o fst) ax_counts) 1 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 773 |               THEN TRY (REPEAT_ALL_NEW (etac @{thm exE}) 1)
 | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 774 | THEN rename_tac outer_param_names 1 | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 775 | THEN copy_prems_tac (map snd ax_counts) [] 1 | 
| 40258 
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
 blanchet parents: 
40221diff
changeset | 776 | THEN release_clusters_tac thy ax_counts substs ordered_clusters 1 | 
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 777 | THEN match_tac [prems_imp_false] 1 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 778 | THEN ALLGOALS (fn i => | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 779 |                        rtac @{thm Meson.skolem_COMBK_I} i
 | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 780 | THEN rotate_tac (rotation_for_subgoal i) i | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 781 | THEN assume_tac i))) | 
| 40158 | 782 | handle ERROR _ => | 
| 783 |              error ("Cannot replay Metis proof in Isabelle:\n\
 | |
| 784 | \Error when discharging Skolem assumptions.") | |
| 39964 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 785 | end | 
| 
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
 blanchet parents: 
39958diff
changeset | 786 | |
| 40665 
1a65f0c74827
added "verbose" option to Metis to shut up its warnings if necessary
 blanchet parents: 
40264diff
changeset | 787 | val setup = trace_setup #> verbose_setup | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 788 | |
| 39495 
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
 blanchet parents: diff
changeset | 789 | end; |