321 | prop_of_clause names = |
321 | prop_of_clause names = |
322 let |
322 let |
323 val lits = |
323 val lits = |
324 map (maybe_dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names) |
324 map (maybe_dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names) |
325 in |
325 in |
326 case List.partition (can HOLogic.dest_not) lits of |
326 (case List.partition (can HOLogic.dest_not) lits of |
327 (negs as _ :: _, pos as _ :: _) => |
327 (negs as _ :: _, pos as _ :: _) => |
328 s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos) |
328 s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos) |
329 | _ => fold (curry s_disj) lits @{term False} |
329 | _ => fold (curry s_disj) lits @{term False}) |
330 end |
330 end |
331 |> HOLogic.mk_Trueprop |> close_form |
331 |> HOLogic.mk_Trueprop |> close_form |
332 |
332 |
333 fun maybe_show outer c = |
333 fun maybe_show outer c = |
334 (outer andalso length c = 1 andalso subset (op =) (c, conjs)) ? cons Show |
334 (outer andalso length c = 1 andalso subset (op =) (c, conjs)) ? cons Show |
378 isar_steps outer (SOME l) (step :: accum) infs |
378 isar_steps outer (SOME l) (step :: accum) infs |
379 end |
379 end |
380 and isar_proof outer fix assms lems infs = |
380 and isar_proof outer fix assms lems infs = |
381 Proof (fix, assms, lems @ isar_steps outer NONE [] infs) |
381 Proof (fix, assms, lems @ isar_steps outer NONE [] infs) |
382 |
382 |
383 val isar_proof_of_direct_proof = isar_proof true params assms lems |
|
384 |
|
385 (* 60 seconds seems like a good interpreation of "no timeout" *) |
383 (* 60 seconds seems like a good interpreation of "no timeout" *) |
386 val preplay_timeout = preplay_timeout |> the_default (seconds 60.0) |
384 val preplay_timeout = preplay_timeout |> the_default (seconds 60.0) |
387 |
385 |
388 val clean_up_labels_in_proof = |
|
389 chain_direct_proof |
|
390 #> kill_useless_labels_in_proof |
|
391 #> relabel_proof |
|
392 val (preplay_interface as {overall_preplay_stats, ...}, isar_proof) = |
386 val (preplay_interface as {overall_preplay_stats, ...}, isar_proof) = |
393 refute_graph |
387 refute_graph |
394 (* |
388 (* |
395 |> tap (tracing o prefix "Refute graph: " o string_of_refute_graph) |
389 |> tap (tracing o prefix "Refute graph: " o string_of_refute_graph) |
396 *) |
390 *) |
397 |> redirect_graph axioms tainted bot |
391 |> redirect_graph axioms tainted bot |
398 (* |
392 (* |
399 |> tap (tracing o prefix "Direct proof: " o string_of_direct_proof) |
393 |> tap (tracing o prefix "Direct proof: " o string_of_direct_proof) |
400 *) |
394 *) |
401 |> isar_proof_of_direct_proof |
395 |> isar_proof true params assms lems |
402 |> postprocess_remove_unreferenced_steps I |
396 |> postprocess_remove_unreferenced_steps I |
403 |> relabel_proof_canonically |
397 |> relabel_proof_canonically |
404 |> `(proof_preplay_interface debug ctxt metis_type_enc metis_lam_trans do_preplay |
398 |> `(proof_preplay_interface debug ctxt metis_type_enc metis_lam_trans do_preplay |
405 preplay_timeout) |
399 preplay_timeout) |
|
400 |
406 val ((preplay_time, preplay_fail), isar_proof) = |
401 val ((preplay_time, preplay_fail), isar_proof) = |
407 isar_proof |
402 isar_proof |
408 |> compress_proof (if isar_proofs = SOME true then isar_compress else 1000.0) |
403 |> compress_proof (if isar_proofs = SOME true then isar_compress else 1000.0) |
409 preplay_interface |
404 preplay_interface |
410 |> isar_try0 ? try0 preplay_timeout preplay_interface |
405 |> isar_try0 ? try0 preplay_timeout preplay_interface |
411 |> postprocess_remove_unreferenced_steps (isar_try0 ? min_deps_of_step preplay_interface) |
406 |> postprocess_remove_unreferenced_steps (isar_try0 ? min_deps_of_step preplay_interface) |
412 |> `overall_preplay_stats |
407 |> `overall_preplay_stats |
413 ||> clean_up_labels_in_proof |
408 ||> (chain_direct_proof |
|
409 #> kill_useless_labels_in_proof |
|
410 #> relabel_proof) |
|
411 |
414 val isar_text = |
412 val isar_text = |
415 string_of_proof ctxt metis_type_enc metis_lam_trans subgoal subgoal_count isar_proof |
413 string_of_proof ctxt metis_type_enc metis_lam_trans subgoal subgoal_count isar_proof |
416 in |
414 in |
417 case isar_text of |
415 (case isar_text of |
418 "" => |
416 "" => |
419 if isar_proofs = SOME true then |
417 if isar_proofs = SOME true then |
420 "\nNo structured proof available (proof too simple)." |
418 "\nNo structured proof available (proof too simple)." |
421 else |
419 else |
422 "" |
420 "" |
435 else |
433 else |
436 []) |
434 []) |
437 in |
435 in |
438 "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^ |
436 "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^ |
439 Active.sendback_markup [Markup.padding_command] isar_text |
437 Active.sendback_markup [Markup.padding_command] isar_text |
440 end |
438 end) |
441 end |
439 end |
442 val isar_proof = |
440 val isar_proof = |
443 if debug then |
441 if debug then |
444 isar_proof_of () |
442 isar_proof_of () |
445 else case try isar_proof_of () of |
443 else |
446 SOME s => s |
444 (case try isar_proof_of () of |
447 | NONE => if isar_proofs = SOME true then |
445 SOME s => s |
448 "\nWarning: The Isar proof construction failed." |
446 | NONE => |
449 else |
447 if isar_proofs = SOME true then "\nWarning: The Isar proof construction failed." else "") |
450 "" |
|
451 in one_line_proof ^ isar_proof end |
448 in one_line_proof ^ isar_proof end |
452 |
449 |
453 fun isar_proof_would_be_a_good_idea preplay = |
450 fun isar_proof_would_be_a_good_idea preplay = |
454 case preplay of |
451 (case preplay of |
455 Played (reconstr, _) => reconstr = SMT |
452 Played (reconstr, _) => reconstr = SMT |
456 | Trust_Playable _ => false |
453 | Trust_Playable _ => false |
457 | Failed_to_Play _ => true |
454 | Failed_to_Play _ => true) |
458 |
455 |
459 fun proof_text ctxt isar_proofs isar_params num_chained |
456 fun proof_text ctxt isar_proofs isar_params num_chained |
460 (one_line_params as (preplay, _, _, _, _, _)) = |
457 (one_line_params as (preplay, _, _, _, _, _)) = |
461 (if isar_proofs = SOME true orelse |
458 (if isar_proofs = SOME true orelse |
462 (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then |
459 (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then |