src/Pure/Isar/obtain.ML
author wenzelm
Wed, 24 Jun 2009 21:28:02 +0200
changeset 31794 71af1fd6a5e4
parent 30770 6976521b4263
child 32111 30e2ffbba718
permissions -rw-r--r--
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
renamed Variable.importT_thms to Variable.importT (again);
     1 (*  Title:      Pure/Isar/obtain.ML
     2     Author:     Markus Wenzel, TU Muenchen
     3 
     4 The 'obtain' and 'guess' language elements -- generalized existence at
     5 the level of proof texts: 'obtain' involves a proof that certain
     6 fixes/assumes may be introduced into the present context; 'guess' is
     7 similar, but derives these elements from the course of reasoning!
     8 
     9   <chain_facts>
    10   obtain x where "A x" <proof> ==
    11 
    12   have "!!thesis. (!!x. A x ==> thesis) ==> thesis"
    13   proof succeed
    14     fix thesis
    15     assume that [intro?]: "!!x. A x ==> thesis"
    16     <chain_facts>
    17     show thesis
    18       apply (insert that)
    19       <proof>
    20   qed
    21   fix x assm <<obtain_export>> "A x"
    22 
    23 
    24   <chain_facts>
    25   guess x <proof body> <proof end> ==
    26 
    27   {
    28     fix thesis
    29     <chain_facts> have "PROP ?guess"
    30       apply magic      -- {* turns goal into "thesis ==> #thesis" *}
    31       <proof body>
    32       apply_end magic  -- {* turns final "(!!x. P x ==> thesis) ==> #thesis" into
    33         "#((!!x. A x ==> thesis) ==> thesis)" which is a finished goal state *}
    34       <proof end>
    35   }
    36   fix x assm <<obtain_export>> "A x"
    37 *)
    38 
    39 signature OBTAIN =
    40 sig
    41   val thatN: string
    42   val obtain: string -> (binding * string option * mixfix) list ->
    43     (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state
    44   val obtain_i: string -> (binding * typ option * mixfix) list ->
    45     (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state
    46   val result: (Proof.context -> tactic) -> thm list -> Proof.context ->
    47     (cterm list * thm list) * Proof.context
    48   val guess: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
    49   val guess_i: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
    50 end;
    51 
    52 structure Obtain: OBTAIN =
    53 struct
    54 
    55 (** obtain_export **)
    56 
    57 (*
    58   [x, A x]
    59      :
    60      B
    61   --------
    62      B
    63 *)
    64 fun eliminate_term ctxt xs tm =
    65   let
    66     val vs = map (dest_Free o Thm.term_of) xs;
    67     val bads = Term.fold_aterms (fn t as Free v =>
    68       if member (op =) vs v then insert (op aconv) t else I | _ => I) tm [];
    69     val _ = null bads orelse
    70       error ("Result contains obtained parameters: " ^
    71         space_implode " " (map (Syntax.string_of_term ctxt) bads));
    72   in tm end;
    73 
    74 fun eliminate fix_ctxt rule xs As thm =
    75   let
    76     val thy = ProofContext.theory_of fix_ctxt;
    77 
    78     val _ = eliminate_term fix_ctxt xs (Thm.full_prop_of thm);
    79     val _ = ObjectLogic.is_judgment thy (Thm.concl_of thm) orelse
    80       error "Conclusion in obtained context must be object-logic judgment";
    81 
    82     val ((_, [thm']), ctxt') = Variable.import true [thm] fix_ctxt;
    83     val prems = Drule.strip_imp_prems (#prop (Thm.crep_thm thm'));
    84   in
    85     ((Drule.implies_elim_list thm' (map Thm.assume prems)
    86         |> Drule.implies_intr_list (map Drule.norm_hhf_cterm As)
    87         |> Drule.forall_intr_list xs)
    88       COMP rule)
    89     |> Drule.implies_intr_list prems
    90     |> singleton (Variable.export ctxt' fix_ctxt)
    91   end;
    92 
    93 fun obtain_export ctxt rule xs _ As =
    94   (eliminate ctxt rule xs As, eliminate_term ctxt xs);
    95 
    96 
    97 
    98 (** obtain **)
    99 
   100 fun bind_judgment ctxt name =
   101   let
   102     val (bind, ctxt') = ProofContext.bind_fixes [name] ctxt;
   103     val (t as _ $ Free v) = bind (ObjectLogic.fixed_judgment (ProofContext.theory_of ctxt) name);
   104   in ((v, t), ctxt') end;
   105 
   106 val thatN = "that";
   107 
   108 local
   109 
   110 fun gen_obtain prep_att prep_vars prep_propp
   111     name raw_vars raw_asms int state =
   112   let
   113     val _ = Proof.assert_forward_or_chain state;
   114     val thy = Proof.theory_of state;
   115     val cert = Thm.cterm_of thy;
   116     val ctxt = Proof.context_of state;
   117     val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
   118 
   119     (*obtain vars*)
   120     val (vars, vars_ctxt) = prep_vars raw_vars ctxt;
   121     val (_, fix_ctxt) = vars_ctxt |> ProofContext.add_fixes vars;
   122     val xs = map (Name.of_binding o #1) vars;
   123 
   124     (*obtain asms*)
   125     val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
   126     val asm_props = maps (map fst) proppss;
   127     val asms = map fst (Attrib.map_specs (prep_att thy) raw_asms) ~~ proppss;
   128 
   129     val _ = Variable.warn_extra_tfrees fix_ctxt asms_ctxt;
   130 
   131     (*obtain statements*)
   132     val thesisN = Name.variant xs AutoBind.thesisN;
   133     val (thesis_var, thesis) = #1 (bind_judgment fix_ctxt thesisN);
   134 
   135     val asm_frees = fold Term.add_frees asm_props [];
   136     val parms = xs |> map (fn x =>
   137       let val x' = ProofContext.get_skolem fix_ctxt x
   138       in (x', the_default propT (AList.lookup (op =) asm_frees x')) end);
   139 
   140     val that_name = if name = "" then thatN else name;
   141     val that_prop =
   142       Term.list_all_free (parms, Logic.list_implies (asm_props, thesis))
   143       |> Library.curry Logic.list_rename_params xs;
   144     val obtain_prop =
   145       Logic.list_rename_params ([AutoBind.thesisN],
   146         Term.list_all_free ([thesis_var], Logic.mk_implies (that_prop, thesis)));
   147 
   148     fun after_qed _ =
   149       Proof.local_qed (NONE, false)
   150       #> `Proof.the_fact #-> (fn rule =>
   151         Proof.fix_i vars
   152         #> Proof.assm_i (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms);
   153   in
   154     state
   155     |> Proof.enter_forward
   156     |> Proof.have_i NONE (K I) [(Thm.empty_binding, [(obtain_prop, [])])] int
   157     |> Proof.proof (SOME Method.succeed_text) |> Seq.hd
   158     |> Proof.fix_i [(Binding.name thesisN, NONE, NoSyn)]
   159     |> Proof.assume_i
   160       [((Binding.name that_name, [ContextRules.intro_query NONE]), [(that_prop, [])])]
   161     |> `Proof.the_facts
   162     ||> Proof.chain_facts chain_facts
   163     ||> Proof.show_i NONE after_qed [(Thm.empty_binding, [(thesis, [])])] false
   164     |-> Proof.refine_insert
   165   end;
   166 
   167 in
   168 
   169 val obtain = gen_obtain Attrib.attribute ProofContext.read_vars ProofContext.read_propp;
   170 val obtain_i = gen_obtain (K I) ProofContext.cert_vars ProofContext.cert_propp;
   171 
   172 end;
   173 
   174 
   175 
   176 (** tactical result **)
   177 
   178 fun check_result ctxt thesis th =
   179   (case Thm.prems_of th of
   180     [prem] =>
   181       if Thm.concl_of th aconv thesis andalso
   182         Logic.strip_assums_concl prem aconv thesis then th
   183       else error ("Guessed a different clause:\n" ^ ProofContext.string_of_thm ctxt th)
   184   | [] => error "Goal solved -- nothing guessed."
   185   | _ => error ("Guess split into several cases:\n" ^ ProofContext.string_of_thm ctxt th));
   186 
   187 fun result tac facts ctxt =
   188   let
   189     val thy = ProofContext.theory_of ctxt;
   190     val cert = Thm.cterm_of thy;
   191 
   192     val ((thesis_var, thesis), thesis_ctxt) = bind_judgment ctxt AutoBind.thesisN;
   193     val rule =
   194       (case SINGLE (Method.insert_tac facts 1 THEN tac thesis_ctxt) (Goal.init (cert thesis)) of
   195         NONE => raise THM ("Obtain.result: tactic failed", 0, facts)
   196       | SOME th => check_result ctxt thesis (MetaSimplifier.norm_hhf (Goal.conclude th)));
   197 
   198     val closed_rule = Thm.forall_intr (cert (Free thesis_var)) rule;
   199     val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt;
   200     val obtain_rule = Thm.forall_elim (cert (Logic.varify (Free thesis_var))) rule';
   201     val ((params, stmt), fix_ctxt) = Variable.focus (Thm.cprem_of obtain_rule 1) ctxt';
   202     val (prems, ctxt'') =
   203       Assumption.add_assms (obtain_export fix_ctxt obtain_rule params)
   204         (Drule.strip_imp_prems stmt) fix_ctxt;
   205   in ((params, prems), ctxt'') end;
   206 
   207 
   208 
   209 (** guess **)
   210 
   211 local
   212 
   213 fun unify_params vars thesis_var raw_rule ctxt =
   214   let
   215     val thy = ProofContext.theory_of ctxt;
   216     val certT = Thm.ctyp_of thy;
   217     val cert = Thm.cterm_of thy;
   218     val string_of_typ = Syntax.string_of_typ ctxt;
   219     val string_of_term = setmp show_types true (Syntax.string_of_term ctxt);
   220 
   221     fun err msg th = error (msg ^ ":\n" ^ ProofContext.string_of_thm ctxt th);
   222 
   223     val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1;
   224     val rule = Thm.incr_indexes (maxidx + 1) raw_rule;
   225 
   226     val params = RuleCases.strip_params (Logic.nth_prem (1, Thm.prop_of rule));
   227     val m = length vars;
   228     val n = length params;
   229     val _ = m <= n orelse err "More variables than parameters in obtained rule" rule;
   230 
   231     fun unify ((x, T), (y, U)) (tyenv, max) = Sign.typ_unify thy (T, U) (tyenv, max)
   232       handle Type.TUNIFY =>
   233         err ("Failed to unify variable " ^
   234           string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^
   235           string_of_term (Syntax.mark_boundT (y, Envir.norm_type tyenv U)) ^ " in") rule;
   236     val (tyenv, _) = fold unify (map #1 vars ~~ Library.take (m, params))
   237       (Vartab.empty, Int.max (maxidx, Thm.maxidx_of rule));
   238     val norm_type = Envir.norm_type tyenv;
   239 
   240     val xs = map (apsnd norm_type o fst) vars;
   241     val ys = map (apsnd norm_type) (Library.drop (m, params));
   242     val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys;
   243     val terms = map (Drule.mk_term o cert o Free) (xs @ ys');
   244 
   245     val instT =
   246       fold (Term.add_tvarsT o #2) params []
   247       |> map (TVar #> (fn T => (certT T, certT (norm_type T))));
   248     val closed_rule = rule
   249       |> Thm.forall_intr (cert (Free thesis_var))
   250       |> Thm.instantiate (instT, []);
   251 
   252     val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt;
   253     val vars' =
   254       map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~
   255       (map snd vars @ replicate (length ys) NoSyn);
   256     val rule'' = Thm.forall_elim (cert (Logic.varify (Free thesis_var))) rule';
   257   in ((vars', rule''), ctxt') end;
   258 
   259 fun inferred_type (binding, _, mx) ctxt =
   260   let
   261     val x = Name.of_binding binding;
   262     val (T, ctxt') = ProofContext.inferred_param x ctxt
   263   in ((x, T, mx), ctxt') end;
   264 
   265 fun polymorphic ctxt vars =
   266   let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars))
   267   in map2 (fn (x, _, mx) => fn T => ((x, T), mx)) vars Ts end;
   268 
   269 fun gen_guess prep_vars raw_vars int state =
   270   let
   271     val _ = Proof.assert_forward_or_chain state;
   272     val thy = Proof.theory_of state;
   273     val cert = Thm.cterm_of thy;
   274     val ctxt = Proof.context_of state;
   275     val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
   276 
   277     val (thesis_var, thesis) = #1 (bind_judgment ctxt AutoBind.thesisN);
   278     val vars = ctxt |> prep_vars raw_vars |-> fold_map inferred_type |> fst |> polymorphic ctxt;
   279 
   280     fun guess_context raw_rule state' =
   281       let
   282         val ((parms, rule), ctxt') =
   283           unify_params vars thesis_var raw_rule (Proof.context_of state');
   284         val (bind, _) = ProofContext.bind_fixes (map (#1 o #1) parms) ctxt';
   285         val ts = map (bind o Free o #1) parms;
   286         val ps = map dest_Free ts;
   287         val asms =
   288           Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule))
   289           |> map (fn asm => (Term.betapplys (Term.list_abs (ps, asm), ts), []));
   290         val _ = not (null asms) orelse error "Trivial result -- nothing guessed";
   291       in
   292         state'
   293         |> Proof.map_context (K ctxt')
   294         |> Proof.fix_i (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms)
   295         |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm_i
   296           (obtain_export fix_ctxt rule (map cert ts)) [(Thm.empty_binding, asms)])
   297         |> Proof.bind_terms AutoBind.no_facts
   298       end;
   299 
   300     val goal = Var (("guess", 0), propT);
   301     fun print_result ctxt' (k, [(s, [_, th])]) =
   302       ProofDisplay.print_results int ctxt' (k, [(s, [th])]);
   303     val before_qed = SOME (Method.primitive_text (Goal.conclude #> MetaSimplifier.norm_hhf #>
   304         (fn th => Goal.protect (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th))));
   305     fun after_qed [[_, res]] =
   306       Proof.end_block #> guess_context (check_result ctxt thesis res);
   307   in
   308     state
   309     |> Proof.enter_forward
   310     |> Proof.begin_block
   311     |> Proof.fix_i [(Binding.name AutoBind.thesisN, NONE, NoSyn)]
   312     |> Proof.chain_facts chain_facts
   313     |> Proof.local_goal print_result (K I) (apsnd (rpair I))
   314       "guess" before_qed after_qed [(Thm.empty_binding, [Logic.mk_term goal, goal])]
   315     |> Proof.refine (Method.primitive_text (K (Goal.init (cert thesis)))) |> Seq.hd
   316   end;
   317 
   318 in
   319 
   320 val guess = gen_guess ProofContext.read_vars;
   321 val guess_i = gen_guess ProofContext.cert_vars;
   322 
   323 end;
   324 
   325 end;