src/Pure/goal.ML
author wenzelm
Sat, 16 Apr 2011 22:21:34 +0200
changeset 43241 244911efd275
parent 43231 da8817d01e7c
child 43242 5900f06b4198
permissions -rw-r--r--
refined PARALLEL_GOALS;
     1 (*  Title:      Pure/goal.ML
     2     Author:     Makarius
     3 
     4 Goals in tactical theorem proving.
     5 *)
     6 
     7 signature BASIC_GOAL =
     8 sig
     9   val parallel_proofs: int Unsynchronized.ref
    10   val parallel_proofs_threshold: int Unsynchronized.ref
    11   val SELECT_GOAL: tactic -> int -> tactic
    12   val CONJUNCTS: tactic -> int -> tactic
    13   val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic
    14   val PARALLEL_CHOICE: tactic list -> tactic
    15   val PARALLEL_GOALS: tactic -> tactic
    16 end;
    17 
    18 signature GOAL =
    19 sig
    20   include BASIC_GOAL
    21   val init: cterm -> thm
    22   val protect: thm -> thm
    23   val conclude: thm -> thm
    24   val check_finished: Proof.context -> thm -> unit
    25   val finish: Proof.context -> thm -> thm
    26   val norm_result: thm -> thm
    27   val fork_name: string -> (unit -> 'a) -> 'a future
    28   val fork: (unit -> 'a) -> 'a future
    29   val future_enabled: unit -> bool
    30   val local_future_enabled: unit -> bool
    31   val future_result: Proof.context -> thm future -> term -> thm
    32   val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
    33   val prove_multi: Proof.context -> string list -> term list -> term list ->
    34     ({prems: thm list, context: Proof.context} -> tactic) -> thm list
    35   val prove_future: Proof.context -> string list -> term list -> term ->
    36     ({prems: thm list, context: Proof.context} -> tactic) -> thm
    37   val prove: Proof.context -> string list -> term list -> term ->
    38     ({prems: thm list, context: Proof.context} -> tactic) -> thm
    39   val prove_global: theory -> string list -> term list -> term ->
    40     ({prems: thm list, context: Proof.context} -> tactic) -> thm
    41   val extract: int -> int -> thm -> thm Seq.seq
    42   val retrofit: int -> int -> thm -> thm -> thm Seq.seq
    43   val conjunction_tac: int -> tactic
    44   val precise_conjunction_tac: int -> int -> tactic
    45   val recover_conjunction_tac: tactic
    46   val norm_hhf_tac: int -> tactic
    47   val compose_hhf_tac: thm -> int -> tactic
    48   val assume_rule_tac: Proof.context -> int -> tactic
    49 end;
    50 
    51 structure Goal: GOAL =
    52 struct
    53 
    54 (** goals **)
    55 
    56 (*
    57   -------- (init)
    58   C ==> #C
    59 *)
    60 val init =
    61   let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI))
    62   in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end;
    63 
    64 (*
    65    C
    66   --- (protect)
    67   #C
    68 *)
    69 fun protect th = Drule.comp_no_flatten (th, 0) 1 Drule.protectI;
    70 
    71 (*
    72   A ==> ... ==> #C
    73   ---------------- (conclude)
    74   A ==> ... ==> C
    75 *)
    76 fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD;
    77 
    78 (*
    79   #C
    80   --- (finish)
    81    C
    82 *)
    83 fun check_finished ctxt th =
    84   (case Thm.nprems_of th of
    85     0 => ()
    86   | n => raise THM ("Proof failed.\n" ^
    87       Pretty.string_of (Pretty.chunks
    88         (Goal_Display.pretty_goals
    89           (ctxt
    90             |> Config.put Goal_Display.goals_limit n
    91             |> Config.put Goal_Display.show_main_goal true) th)) ^
    92       "\n" ^ string_of_int n ^ " unsolved goal(s)!", 0, [th]));
    93 
    94 fun finish ctxt th = (check_finished ctxt th; conclude th);
    95 
    96 
    97 
    98 (** results **)
    99 
   100 (* normal form *)
   101 
   102 val norm_result =
   103   Drule.flexflex_unique
   104   #> Raw_Simplifier.norm_hhf_protect
   105   #> Thm.strip_shyps
   106   #> Drule.zero_var_indexes;
   107 
   108 
   109 (* forked proofs *)
   110 
   111 val forked_proofs = Synchronized.var "forked_proofs" 0;
   112 
   113 fun forked i =
   114   Synchronized.change forked_proofs (fn m =>
   115     let
   116       val n = m + i;
   117       val _ =
   118         Multithreading.tracing 2 (fn () =>
   119           ("PROOFS " ^ Time.toString (Time.now ()) ^ ": " ^ string_of_int n));
   120     in n end);
   121 
   122 fun fork_name name e =
   123   uninterruptible (fn _ => fn () =>
   124     let
   125       val _ = forked 1;
   126       val future =
   127         singleton (Future.forks {name = name, group = NONE, deps = [], pri = ~1})
   128           (fn () =>
   129             (*interruptible*)
   130             Exn.release
   131               (Exn.capture Future.status e before forked ~1
   132                 handle exn => (forked ~1; reraise exn)));
   133     in future end) ();
   134 
   135 fun fork e = fork_name "Goal.fork" e;
   136 
   137 
   138 (* scheduling parameters *)
   139 
   140 val parallel_proofs = Unsynchronized.ref 1;
   141 val parallel_proofs_threshold = Unsynchronized.ref 50;
   142 
   143 fun future_enabled () =
   144   Multithreading.enabled () andalso ! parallel_proofs >= 1 andalso
   145   is_some (Future.worker_task ());
   146 
   147 fun local_future_enabled () =
   148   future_enabled () andalso ! parallel_proofs >= 2 andalso
   149   Synchronized.value forked_proofs <
   150     ! parallel_proofs_threshold * Multithreading.max_threads_value ();
   151 
   152 
   153 (* future_result *)
   154 
   155 fun future_result ctxt result prop =
   156   let
   157     val thy = Proof_Context.theory_of ctxt;
   158     val _ = Context.reject_draft thy;
   159     val cert = Thm.cterm_of thy;
   160     val certT = Thm.ctyp_of thy;
   161 
   162     val assms = Assumption.all_assms_of ctxt;
   163     val As = map Thm.term_of assms;
   164 
   165     val xs = map Free (fold Term.add_frees (prop :: As) []);
   166     val fixes = map cert xs;
   167 
   168     val tfrees = fold Term.add_tfrees (prop :: As) [];
   169     val instT = map (fn (a, S) => (certT (TVar ((a, 0), S)), certT (TFree (a, S)))) tfrees;
   170 
   171     val global_prop =
   172       cert (Term.map_types Logic.varifyT_global
   173         (fold_rev Logic.all xs (Logic.list_implies (As, prop))))
   174       |> Thm.weaken_sorts (Variable.sorts_of ctxt);
   175     val global_result = result |> Future.map
   176       (Drule.flexflex_unique #>
   177         Thm.adjust_maxidx_thm ~1 #>
   178         Drule.implies_intr_list assms #>
   179         Drule.forall_intr_list fixes #>
   180         Thm.generalize (map #1 tfrees, []) 0 #>
   181         Thm.strip_shyps);
   182     val local_result =
   183       Thm.future global_result global_prop
   184       |> Thm.instantiate (instT, [])
   185       |> Drule.forall_elim_list fixes
   186       |> fold (Thm.elim_implies o Thm.assume) assms;
   187   in local_result end;
   188 
   189 
   190 
   191 (** tactical theorem proving **)
   192 
   193 (* prove_internal -- minimal checks, no normalization of result! *)
   194 
   195 fun prove_internal casms cprop tac =
   196   (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
   197     SOME th => Drule.implies_intr_list casms
   198       (finish (Syntax.init_pretty_global (Thm.theory_of_thm th)) th)
   199   | NONE => error "Tactic failed");
   200 
   201 
   202 (* prove_common etc. *)
   203 
   204 fun prove_common immediate ctxt xs asms props tac =
   205   let
   206     val thy = Proof_Context.theory_of ctxt;
   207     val string_of_term = Syntax.string_of_term ctxt;
   208 
   209     val pos = Position.thread_data ();
   210     fun err msg = cat_error msg
   211       ("The error(s) above occurred for the goal statement:\n" ^
   212         string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)) ^
   213         (case Position.str_of pos of "" => "" | s => "\n" ^ s));
   214 
   215     fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t))
   216       handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;
   217     val casms = map cert_safe asms;
   218     val cprops = map cert_safe props;
   219 
   220     val (prems, ctxt') = ctxt
   221       |> Variable.add_fixes_direct xs
   222       |> fold Variable.declare_term (asms @ props)
   223       |> Assumption.add_assumes casms
   224       ||> Variable.set_body true;
   225     val sorts = Variable.sorts_of ctxt';
   226 
   227     val stmt = Thm.weaken_sorts sorts (Conjunction.mk_conjunction_balanced cprops);
   228 
   229     fun result () =
   230       (case SINGLE (tac {prems = prems, context = ctxt'}) (init stmt) of
   231         NONE => err "Tactic failed"
   232       | SOME st =>
   233           let val res = finish ctxt' st handle THM (msg, _, _) => err msg in
   234             if Unify.matches_list thy [Thm.term_of stmt] [Thm.prop_of res]
   235             then Thm.check_shyps sorts res
   236             else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res))
   237           end);
   238     val res =
   239       if immediate orelse #maxidx (Thm.rep_cterm stmt) >= 0 orelse not (future_enabled ())
   240       then result ()
   241       else future_result ctxt' (fork result) (Thm.term_of stmt);
   242   in
   243     Conjunction.elim_balanced (length props) res
   244     |> map (Assumption.export false ctxt' ctxt)
   245     |> Variable.export ctxt' ctxt
   246     |> map Drule.zero_var_indexes
   247   end;
   248 
   249 val prove_multi = prove_common true;
   250 
   251 fun prove_future ctxt xs asms prop tac = hd (prove_common false ctxt xs asms [prop] tac);
   252 fun prove ctxt xs asms prop tac = hd (prove_common true ctxt xs asms [prop] tac);
   253 
   254 fun prove_global thy xs asms prop tac =
   255   Drule.export_without_context (prove (Proof_Context.init_global thy) xs asms prop tac);
   256 
   257 
   258 
   259 (** goal structure **)
   260 
   261 (* nested goals *)
   262 
   263 fun extract i n st =
   264   (if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty
   265    else if n = 1 then Seq.single (Thm.cprem_of st i)
   266    else
   267      Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1))))
   268   |> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init);
   269 
   270 fun retrofit i n st' st =
   271   (if n = 1 then st
   272    else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n))
   273   |> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i;
   274 
   275 fun SELECT_GOAL tac i st =
   276   if Thm.nprems_of st = 1 andalso i = 1 then tac st
   277   else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st;
   278 
   279 
   280 (* multiple goals *)
   281 
   282 fun precise_conjunction_tac 0 i = eq_assume_tac i
   283   | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i
   284   | precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n));
   285 
   286 val adhoc_conjunction_tac = REPEAT_ALL_NEW
   287   (SUBGOAL (fn (goal, i) =>
   288     if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i
   289     else no_tac));
   290 
   291 val conjunction_tac = SUBGOAL (fn (goal, i) =>
   292   precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE
   293   TRY (adhoc_conjunction_tac i));
   294 
   295 val recover_conjunction_tac = PRIMITIVE (fn th =>
   296   Conjunction.uncurry_balanced (Thm.nprems_of th) th);
   297 
   298 fun PRECISE_CONJUNCTS n tac =
   299   SELECT_GOAL (precise_conjunction_tac n 1
   300     THEN tac
   301     THEN recover_conjunction_tac);
   302 
   303 fun CONJUNCTS tac =
   304   SELECT_GOAL (conjunction_tac 1
   305     THEN tac
   306     THEN recover_conjunction_tac);
   307 
   308 
   309 (* hhf normal form *)
   310 
   311 val norm_hhf_tac =
   312   rtac Drule.asm_rl  (*cheap approximation -- thanks to builtin Logic.flatten_params*)
   313   THEN' SUBGOAL (fn (t, i) =>
   314     if Drule.is_norm_hhf t then all_tac
   315     else rewrite_goal_tac Drule.norm_hhf_eqs i);
   316 
   317 fun compose_hhf_tac th i st =
   318   PRIMSEQ (Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of st i) th, 0) i) st;
   319 
   320 
   321 (* non-atomic goal assumptions *)
   322 
   323 fun non_atomic (Const ("==>", _) $ _ $ _) = true
   324   | non_atomic (Const ("all", _) $ _) = true
   325   | non_atomic _ = false;
   326 
   327 fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) =>
   328   let
   329     val ((_, goal'), ctxt') = Variable.focus goal ctxt;
   330     val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
   331     val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');
   332     val tacs = Rs |> map (fn R =>
   333       Tactic.etac (Raw_Simplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac);
   334   in fold_rev (curry op APPEND') tacs (K no_tac) i end);
   335 
   336 
   337 (* parallel tacticals *)
   338 
   339 (*parallel choice of single results*)
   340 fun PARALLEL_CHOICE tacs st =
   341   (case Par_List.get_some (fn tac => SINGLE tac st) tacs of
   342     NONE => Seq.empty
   343   | SOME st' => Seq.single st');
   344 
   345 (*parallel refinement of non-schematic goal by single results*)
   346 exception FAILED of unit;
   347 fun PARALLEL_GOALS tac =
   348   Thm.adjust_maxidx_thm ~1 #>
   349   (fn st =>
   350     if not (future_enabled ()) orelse Thm.maxidx_of st >= 0 orelse Thm.nprems_of st <= 1
   351     then DETERM tac st
   352     else
   353       let
   354         fun try_tac g =
   355           (case SINGLE tac g of
   356             NONE => raise FAILED ()
   357           | SOME g' => g');
   358 
   359         val goals = Drule.strip_imp_prems (Thm.cprop_of st);
   360         val results = Par_List.map (try_tac o init) goals;
   361       in ALLGOALS (fn i => retrofit i 1 (nth results (i - 1))) st end
   362       handle FAILED () => Seq.empty);
   363 
   364 end;
   365 
   366 structure Basic_Goal: BASIC_GOAL = Goal;
   367 open Basic_Goal;