src/HOL/Tools/Sledgehammer/sledgehammer_isar_compress.ML
changeset 55244 12e1a5d8ee48
parent 55243 66709d41601e
child 55246 e9fba9767d92
equal deleted inserted replaced
55243:66709d41601e 55244:12e1a5d8ee48
    77     (case do_steps steps (rev updates) of
    77     (case do_steps steps (rev updates) of
    78       (steps, []) => steps
    78       (steps, []) => steps
    79     | _ => raise Fail "Sledgehammer_Isar_Compress: update_steps")
    79     | _ => raise Fail "Sledgehammer_Isar_Compress: update_steps")
    80   end
    80   end
    81 
    81 
    82 (* Tries merging the first step into the second step.
    82 (* Tries merging the first step into the second step. *)
    83    FIXME: Arbitrarily picks the second step's method. *)
    83 fun try_merge (Prove (_, [], lbl1, _, [], ((lfs1, gfs1), meths1)))
    84 fun try_merge (Prove (_, [], lbl1, _, [], ((lfs1, gfs1), _)))
    84       (Prove (qs2, fix, lbl2, t, subproofs, ((lfs2, gfs2), meths2))) =
    85       (Prove (qs2, fix, lbl2, t, subproofs, ((lfs2, gfs2), methss2))) =
       
    86     let
    85     let
    87       val lfs = remove (op =) lbl1 lfs2 |> union (op =) lfs1
    86       val lfs = remove (op =) lbl1 lfs2 |> union (op =) lfs1
    88       val gfs = union (op =) gfs1 gfs2
    87       val gfs = union (op =) gfs1 gfs2
    89     in
    88     in
    90       SOME (Prove (qs2, fix, lbl2, t, subproofs, ((lfs, gfs), methss2)))
    89       SOME (Prove (qs2, fix, lbl2, t, subproofs, ((lfs, gfs), meths2)))
    91     end
    90     end
    92   | try_merge _ _ = NONE
    91   | try_merge _ _ = NONE
    93 
    92 
    94 val compress_degree = 2
    93 val compress_degree = 2
    95 val merge_timeout_slack = 1.2
    94 val merge_timeout_slack = 1.2
   134         in
   133         in
   135           (get_successors, replace_successor)
   134           (get_successors, replace_successor)
   136         end
   135         end
   137 
   136 
   138       (* elimination of trivial, one-step subproofs *)
   137       (* elimination of trivial, one-step subproofs *)
   139       fun elim_subproofs' time qs fix l t lfs gfs (methss as (meth :: _) :: _) subs nontriv_subs =
   138       fun elim_subproofs' time qs fix l t lfs gfs (meths as meth :: _) subs nontriv_subs =
   140         if null subs orelse not (compress_further ()) then
   139         if null subs orelse not (compress_further ()) then
   141           (set_preplay_outcome l meth (Played time);
   140           (set_preplay_outcome l meth (Played time);
   142            Prove (qs, fix, l, t, List.revAppend (nontriv_subs, subs), ((lfs, gfs), methss)))
   141            Prove (qs, fix, l, t, List.revAppend (nontriv_subs, subs), ((lfs, gfs), meths)))
   143         else
   142         else
   144           (case subs of
   143           (case subs of
   145             (sub as Proof (_, assms, sub_steps)) :: subs =>
   144             (sub as Proof (_, assms, sub_steps)) :: subs =>
   146             (let
   145             (let
   147               (* trivial subproofs have exactly one "Prove" step *)
   146               (* trivial subproofs have exactly one "Prove" step *)
   148               val SOME (Prove (_, [], l', _, [], ((lfs', gfs'), (meth' :: _) :: _))) =
   147               val SOME (Prove (_, [], l', _, [], ((lfs', gfs'), meth' :: _))) =
   149                 try the_single sub_steps
   148                 try the_single sub_steps
   150 
   149 
   151               (* only touch proofs that can be preplayed sucessfully *)
   150               (* only touch proofs that can be preplayed sucessfully *)
   152               val Played time' = Lazy.force (preplay_outcome l' meth')
   151               val Played time' = Lazy.force (preplay_outcome l' meth')
   153 
   152 
   154               (* merge steps *)
   153               (* merge steps *)
   155               val subs'' = subs @ nontriv_subs
   154               val subs'' = subs @ nontriv_subs
   156               val lfs'' = union (op =) lfs (subtract (op =) (map fst assms) lfs')
   155               val lfs'' = union (op =) lfs (subtract (op =) (map fst assms) lfs')
   157               val gfs'' = union (op =) gfs' gfs
   156               val gfs'' = union (op =) gfs' gfs
   158               val by = ((lfs'', gfs''), methss(*FIXME*))
   157               val by = ((lfs'', gfs''), meths(*FIXME*))
   159               val step'' = Prove (qs, fix, l, t, subs'', by)
   158               val step'' = Prove (qs, fix, l, t, subs'', by)
   160 
   159 
   161               (* check if the modified step can be preplayed fast enough *)
   160               (* check if the modified step can be preplayed fast enough *)
   162               val timeout = time_mult merge_timeout_slack (Time.+(time, time'))
   161               val timeout = time_mult merge_timeout_slack (Time.+(time, time'))
   163               val Played time'' = preplay_quietly timeout step''
   162               val Played time'' = preplay_quietly timeout step''
   164             in
   163             in
   165               decrement_step_count (); (* l' successfully eliminated! *)
   164               decrement_step_count (); (* l' successfully eliminated! *)
   166               map (replace_successor l' [l]) lfs';
   165               map (replace_successor l' [l]) lfs';
   167               elim_subproofs' time'' qs fix l t lfs'' gfs'' methss subs nontriv_subs
   166               elim_subproofs' time'' qs fix l t lfs'' gfs'' meths subs nontriv_subs
   168             end
   167             end
   169             handle Bind =>
   168             handle Bind =>
   170             elim_subproofs' time qs fix l t lfs gfs methss subs (sub :: nontriv_subs))
   169             elim_subproofs' time qs fix l t lfs gfs meths subs (sub :: nontriv_subs))
   171           | _ => raise Fail "Sledgehammer_Isar_Compress: elim_subproofs'")
   170           | _ => raise Fail "Sledgehammer_Isar_Compress: elim_subproofs'")
   172 
   171 
   173       fun elim_subproofs (step as Let _) = step
   172       fun elim_subproofs (step as Let _) = step
   174         | elim_subproofs (step as Prove (qs, fix, l, t, subproofs,
   173         | elim_subproofs (step as Prove (qs, fix, l, t, subproofs,
   175             ((lfs, gfs), methss as (meth :: _) :: _))) =
   174             ((lfs, gfs), meths as meth :: _))) =
   176           if subproofs = [] then
   175           if subproofs = [] then
   177             step
   176             step
   178           else
   177           else
   179             (case Lazy.force (preplay_outcome l meth) of
   178             (case Lazy.force (preplay_outcome l meth) of
   180               Played time => elim_subproofs' time qs fix l t lfs gfs methss subproofs []
   179               Played time => elim_subproofs' time qs fix l t lfs gfs meths subproofs []
   181             | _ => step)
   180             | _ => step)
   182 
   181 
   183       (** top_level compression: eliminate steps by merging them into their successors **)
   182       (** top_level compression: eliminate steps by merging them into their successors **)
   184       fun compress_top_level steps =
   183       fun compress_top_level steps =
   185         let
   184         let
   209                |> fold_index add_cand) []
   208                |> fold_index add_cand) []
   210             end
   209             end
   211 
   210 
   212           fun try_eliminate (i, l, _) succ_lbls steps =
   211           fun try_eliminate (i, l, _) succ_lbls steps =
   213             let
   212             let
   214               val ((cand as Prove (_, _, l, _, _, ((lfs, _), (meth :: _) :: _))) :: steps') =
   213               val ((cand as Prove (_, _, l, _, _, ((lfs, _), meth :: _))) :: steps') = drop i steps
   215                 drop i steps
       
   216 
   214 
   217               val succs = collect_successors steps' succ_lbls
   215               val succs = collect_successors steps' succ_lbls
   218               val succ_meths = map (hd o hd o snd o the o byline_of_isar_step) succs
   216               val succ_meths = map (hd o snd o the o byline_of_isar_step) succs
   219 
   217 
   220               (* only touch steps that can be preplayed successfully *)
   218               (* only touch steps that can be preplayed successfully *)
   221               val Played time = Lazy.force (preplay_outcome l meth)
   219               val Played time = Lazy.force (preplay_outcome l meth)
   222 
   220 
   223               val succs' = map (try_merge cand #> the) succs
   221               val succs' = map (try_merge cand #> the) succs