src/Pure/Isar/obtain.ML
changeset 30223 24d975352879
parent 30211 556d1810cdad
child 30242 aea5d7fa7ef5
equal deleted inserted replaced
30222:4102bbf2af21 30223:24d975352879
   117     val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
   117     val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
   118 
   118 
   119     (*obtain vars*)
   119     (*obtain vars*)
   120     val (vars, vars_ctxt) = prep_vars raw_vars ctxt;
   120     val (vars, vars_ctxt) = prep_vars raw_vars ctxt;
   121     val (_, fix_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars;
   121     val (_, fix_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars;
   122     val xs = map (Binding.base_name o #1) vars;
   122     val xs = map (Binding.name_of o #1) vars;
   123 
   123 
   124     (*obtain asms*)
   124     (*obtain asms*)
   125     val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
   125     val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
   126     val asm_props = maps (map fst) proppss;
   126     val asm_props = maps (map fst) proppss;
   127     val asms = map fst (Attrib.map_specs (prep_att thy) raw_asms) ~~ proppss;
   127     val asms = map fst (Attrib.map_specs (prep_att thy) raw_asms) ~~ proppss;
   256     val rule'' = Thm.forall_elim (cert (Logic.varify (Free thesis_var))) rule';
   256     val rule'' = Thm.forall_elim (cert (Logic.varify (Free thesis_var))) rule';
   257   in ((vars', rule''), ctxt') end;
   257   in ((vars', rule''), ctxt') end;
   258 
   258 
   259 fun inferred_type (binding, _, mx) ctxt =
   259 fun inferred_type (binding, _, mx) ctxt =
   260   let
   260   let
   261     val x = Binding.base_name binding;
   261     val x = Binding.name_of binding;
   262     val (T, ctxt') = ProofContext.inferred_param x ctxt
   262     val (T, ctxt') = ProofContext.inferred_param x ctxt
   263   in ((x, T, mx), ctxt') end;
   263   in ((x, T, mx), ctxt') end;
   264 
   264 
   265 fun polymorphic ctxt vars =
   265 fun polymorphic ctxt vars =
   266   let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars))
   266   let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars))