src/Pure/Isar/obtain.ML
changeset 18678 dd0c569fa43d
parent 18670 c3f445b92aff
child 18693 8ae076ee5e19
equal deleted inserted replaced
18677:01265301db7f 18678:dd0c569fa43d
    50 struct
    50 struct
    51 
    51 
    52 
    52 
    53 (** obtain_export **)
    53 (** obtain_export **)
    54 
    54 
    55 fun obtain_export state parms rule cprops thm =
    55 fun obtain_export ctxt parms rule cprops thm =
    56   let
    56   let
    57     val {thy, prop, maxidx, ...} = Thm.rep_thm thm;
    57     val {thy, prop, maxidx, ...} = Thm.rep_thm thm;
    58     val cparms = map (Thm.cterm_of thy) parms;
    58     val cparms = map (Thm.cterm_of thy) parms;
    59 
    59 
    60     val thm' = thm
    60     val thm' = thm
    65 
    65 
    66     val concl = Logic.strip_assums_concl prop;
    66     val concl = Logic.strip_assums_concl prop;
    67     val bads = parms inter (Term.term_frees concl);
    67     val bads = parms inter (Term.term_frees concl);
    68   in
    68   in
    69     if not (null bads) then
    69     if not (null bads) then
    70       raise Proof.STATE ("Conclusion contains obtained parameters: " ^
    70       error ("Conclusion contains obtained parameters: " ^
    71         space_implode " " (map (ProofContext.string_of_term (Proof.context_of state)) bads), state)
    71         space_implode " " (map (ProofContext.string_of_term ctxt) bads))
    72     else if not (ObjectLogic.is_judgment thy concl) then
    72     else if not (ObjectLogic.is_judgment thy concl) then
    73       raise Proof.STATE ("Conclusion in obtained context must be object-logic judgments", state)
    73       error "Conclusion in obtained context must be object-logic judgments"
    74     else (Tactic.rtac thm' THEN' RANGE elim_tacs) 1 rule
    74     else (Tactic.rtac thm' THEN' RANGE elim_tacs) 1 rule
    75   end;
    75   end;
    76 
    76 
    77 
    77 
    78 
    78 
    90 
    90 
    91 fun gen_obtain prep_att prep_vars prep_propp raw_vars raw_asms int state =
    91 fun gen_obtain prep_att prep_vars prep_propp raw_vars raw_asms int state =
    92   let
    92   let
    93     val _ = Proof.assert_forward_or_chain state;
    93     val _ = Proof.assert_forward_or_chain state;
    94     val ctxt = Proof.context_of state;
    94     val ctxt = Proof.context_of state;
       
    95     val thy = Proof.theory_of state;
    95     val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
    96     val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
    96 
    97 
    97     (*obtain vars*)
    98     (*obtain vars*)
    98     val (vars, vars_ctxt) = prep_vars (map Syntax.no_syn raw_vars) ctxt;
    99     val (vars, vars_ctxt) = prep_vars (map Syntax.no_syn raw_vars) ctxt;
    99     val (_, fix_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars;
   100     val (_, fix_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars;
   101     val Ts = map #2 vars;
   102     val Ts = map #2 vars;
   102 
   103 
   103     (*obtain asms*)
   104     (*obtain asms*)
   104     val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
   105     val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
   105     val asm_props = List.concat (map (map fst) proppss);
   106     val asm_props = List.concat (map (map fst) proppss);
   106     val asms = map fst (Attrib.map_specs (prep_att (Proof.theory_of state)) raw_asms) ~~ proppss;
   107     val asms = map fst (Attrib.map_specs (prep_att thy) raw_asms) ~~ proppss;
   107 
   108 
   108     val _ = ProofContext.warn_extra_tfrees fix_ctxt asms_ctxt;
   109     val _ = ProofContext.warn_extra_tfrees fix_ctxt asms_ctxt;
   109 
   110 
   110     (*obtain statements*)
   111     (*obtain statements*)
   111     val thesisN = Term.variant xs AutoBind.thesisN;
   112     val thesisN = Term.variant xs AutoBind.thesisN;
   127 
   128 
   128     fun after_qed _ =
   129     fun after_qed _ =
   129       Proof.local_qed (NONE, false)
   130       Proof.local_qed (NONE, false)
   130       #> Seq.map (`Proof.the_fact #-> (fn rule =>
   131       #> Seq.map (`Proof.the_fact #-> (fn rule =>
   131         Proof.fix_i (xs ~~ Ts)
   132         Proof.fix_i (xs ~~ Ts)
   132         #> Proof.assm_i (K (obtain_export state parms rule)) asms));
   133         #> Proof.assm_i (K (obtain_export ctxt parms rule)) asms));
   133   in
   134   in
   134     state
   135     state
   135     |> Proof.enter_forward
   136     |> Proof.enter_forward
   136     |> Proof.have_i NONE (K Seq.single) [(("", []), [(obtain_prop, ([], []))])] int
   137     |> Proof.have_i NONE (K Seq.single) [(("", []), [(obtain_prop, ([], []))])] int
   137     |> Proof.proof (SOME Method.succeed_text) |> Seq.hd
   138     |> Proof.proof (SOME Method.succeed_text) |> Seq.hd
   155 
   156 
   156 (** guess **)
   157 (** guess **)
   157 
   158 
   158 local
   159 local
   159 
   160 
   160 fun match_params state vars rule =
   161 fun match_params ctxt vars rule =
   161   let
   162   let
   162     val ctxt = Proof.context_of state;
   163     val thy = ProofContext.theory_of ctxt;
   163     val thy = Proof.theory_of state;
       
   164     val string_of_typ = ProofContext.string_of_typ ctxt;
   164     val string_of_typ = ProofContext.string_of_typ ctxt;
   165     val string_of_term = setmp show_types true (ProofContext.string_of_term ctxt);
   165     val string_of_term = setmp show_types true (ProofContext.string_of_term ctxt);
   166 
   166 
   167     fun err msg th = raise Proof.STATE (msg ^ ":\n" ^ ProofContext.string_of_thm ctxt th, state);
   167     fun err msg th = error (msg ^ ":\n" ^ ProofContext.string_of_thm ctxt th);
   168 
   168 
   169     val params = RuleCases.strip_params (Logic.nth_prem (1, Thm.prop_of rule));
   169     val params = RuleCases.strip_params (Logic.nth_prem (1, Thm.prop_of rule));
   170     val m = length vars;
   170     val m = length vars;
   171     val n = length params;
   171     val n = length params;
   172     val _ = conditional (m > n)
   172     val _ = conditional (m > n)
   212     fun check_result th =
   212     fun check_result th =
   213       (case Thm.prems_of th of
   213       (case Thm.prems_of th of
   214         [prem] =>
   214         [prem] =>
   215           if Thm.concl_of th aconv thesis andalso
   215           if Thm.concl_of th aconv thesis andalso
   216             Logic.strip_assums_concl prem aconv thesis then ()
   216             Logic.strip_assums_concl prem aconv thesis then ()
   217           else raise Proof.STATE ("Guessed a different clause:\n" ^
   217           else error ("Guessed a different clause:\n" ^ ProofContext.string_of_thm ctxt th)
   218             ProofContext.string_of_thm ctxt th, state)
   218       | [] => error "Goal solved -- nothing guessed."
   219       | [] => raise Proof.STATE ("Goal solved -- nothing guessed.", state)
   219       | _ => error ("Guess split into several cases:\n" ^ ProofContext.string_of_thm ctxt th));
   220       | _ => raise Proof.STATE ("Guess split into several cases:\n" ^
       
   221         ProofContext.string_of_thm ctxt th, state));
       
   222 
   220 
   223     fun guess_context raw_rule =
   221     fun guess_context raw_rule =
   224       let
   222       let
   225         val (parms, rule) = match_params state (map (fn (x, T, _) => (x, T)) vars) raw_rule;
   223         val (parms, rule) = match_params ctxt (map (fn (x, T, _) => (x, T)) vars) raw_rule;
   226         val (bind, _) = ProofContext.bind_fixes (map #1 parms) ctxt;
   224         val (bind, _) = ProofContext.bind_fixes (map #1 parms) ctxt;
   227         val ts = map (bind o Free) parms;
   225         val ts = map (bind o Free) parms;
   228         val ps = map dest_Free ts;
   226         val ps = map dest_Free ts;
   229         val asms =
   227         val asms =
   230           Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule))
   228           Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule))
   231           |> map (fn asm => (Term.betapplys (Term.list_abs (ps, asm), ts), ([], [])));
   229           |> map (fn asm => (Term.betapplys (Term.list_abs (ps, asm), ts), ([], [])));
   232         val _ = conditional (null asms) (fn () =>
   230         val _ = conditional (null asms) (fn () => error "Trivial result -- nothing guessed");
   233           raise Proof.STATE ("Trivial result -- nothing guessed", state));
       
   234       in
   231       in
   235         Proof.fix_i (map (apsnd SOME) parms)
   232         Proof.fix_i (map (apsnd SOME) parms)
   236         #> Proof.assm_i (K (obtain_export state ts rule)) [(("", []), asms)]
   233         #> Proof.assm_i (K (obtain_export ctxt ts rule)) [(("", []), asms)]
   237         #> Proof.add_binds_i AutoBind.no_facts
   234         #> Proof.add_binds_i AutoBind.no_facts
   238       end;
   235       end;
   239 
   236 
   240     val before_qed = SOME (Method.primitive_text (Goal.conclude #> Goal.protect));
   237     val before_qed = SOME (Method.primitive_text (Goal.conclude #> Goal.protect));
   241     fun after_qed [[res]] =
   238     fun after_qed [[res]] =