# HG changeset patch # User blanchet # Date 1676476902 -3600 # Node ID 0506c327381467df3f5418014937fe763a3550be # Parent 40b23105a322b729065987db3c7370733477962e removed rarely used error in Sledgehammer diff -r 40b23105a322 -r 0506c3273814 src/HOL/Tools/ATP/atp_proof.ML --- a/src/HOL/Tools/ATP/atp_proof.ML Wed Feb 15 16:44:52 2023 +0100 +++ b/src/HOL/Tools/ATP/atp_proof.ML Wed Feb 15 17:01:42 2023 +0100 @@ -24,7 +24,6 @@ ProofMissing | ProofIncomplete | ProofUnparsable | - UnsoundProof of bool * string list | TimedOut | Inappropriate | OutOfResources | @@ -143,7 +142,6 @@ ProofMissing | ProofIncomplete | ProofUnparsable | - UnsoundProof of bool * string list | TimedOut | Inappropriate | OutOfResources | @@ -160,21 +158,12 @@ else "" -fun from_lemmas [] = "" - | from_lemmas ss = " from " ^ space_implode " " (Try.serial_commas "and" (map quote ss)) - fun string_of_atp_failure MaybeUnprovable = "Problem maybe unprovable" | string_of_atp_failure Unprovable = "Unprovable problem" | string_of_atp_failure GaveUp = "Gave up" | string_of_atp_failure ProofMissing = "Proof missing" | string_of_atp_failure ProofIncomplete = "Proof incomplete" | string_of_atp_failure ProofUnparsable = "Proof unparsable" - | string_of_atp_failure (UnsoundProof (false, ss)) = - "Derived the lemma \"False\"" ^ from_lemmas ss ^ - ", probably due to the use of an unsound type encoding" - | string_of_atp_failure (UnsoundProof (true, ss)) = - "Derived the lemma \"False\"" ^ from_lemmas ss ^ - ", which could be due to a bug in Sledgehammer or to inconsistent axioms (including \"sorry\"s)" | string_of_atp_failure TimedOut = "Timed out" | string_of_atp_failure Inappropriate = "Problem outside the prover's scope" | string_of_atp_failure OutOfResources = "Out of resources" @@ -184,8 +173,7 @@ | string_of_atp_failure Crashed = "Crashed" | string_of_atp_failure InternalError = "Internal prover error" | string_of_atp_failure (UnknownError s) = - "Prover error" ^ - (if s = "" then " (pass the \"verbose\" option for details)" else ":\n" ^ s) + "Prover error" ^ (if s = "" then " (pass the \"verbose\" option for details)" else ":\n" ^ s) fun extract_delimited (begin_delim, end_delim) output = (case first_field begin_delim output of diff -r 40b23105a322 -r 0506c3273814 src/HOL/Tools/ATP/atp_proof_reconstruct.ML --- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML Wed Feb 15 16:44:52 2023 +0100 +++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML Wed Feb 15 17:01:42 2023 +0100 @@ -46,8 +46,6 @@ val used_facts_in_atp_proof : Proof.context -> (string * stature) list -> string atp_proof -> (string * stature) list - val used_facts_in_unsound_atp_proof : Proof.context -> (string * stature) list -> 'a atp_proof -> - string list option val atp_proof_prefers_lifting : string atp_proof -> bool val is_typed_helper_used_in_atp_proof : string atp_proof -> bool val replace_dependencies_in_line : atp_step_name * atp_step_name list -> ('a, 'b) atp_step -> @@ -519,19 +517,6 @@ fun used_facts_in_atp_proof ctxt facts atp_proof = if null atp_proof then facts else fold (add_fact ctxt facts) atp_proof [] -fun used_facts_in_unsound_atp_proof _ _ [] = NONE - | used_facts_in_unsound_atp_proof ctxt facts atp_proof = - let - val used_facts = used_facts_in_atp_proof ctxt facts atp_proof - val all_global_facts = forall (fn (_, (sc, _)) => sc = Global) used_facts - val axiom_used = is_axiom_used_in_proof (is_some o resolve_conjecture) atp_proof - in - if all_global_facts andalso not axiom_used then - SOME (map fst used_facts) - else - NONE - end - val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix (* overapproximation (good enough) *) diff -r 40b23105a322 -r 0506c3273814 src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML Wed Feb 15 16:44:52 2023 +0100 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML Wed Feb 15 17:01:42 2023 +0100 @@ -237,18 +237,10 @@ (state, subgoal, name, "Running command in " ^ string_of_int (Time.toMilliseconds run_time) ^ " ms")) - val outcome = + val _ = (case outcome of - NONE => - (case used_facts_in_unsound_atp_proof ctxt (map fst facts) atp_proof of - SOME facts => - let - val failure = UnsoundProof (is_type_enc_sound type_enc, sort string_ord facts) - in - if debug then (warning (string_of_atp_failure failure); NONE) else SOME failure - end - | NONE => (found_something name; NONE)) - | _ => outcome) + NONE => found_something name + | _ => ()) in (atp_problem_data, (output, run_time, facts, atp_problem, tstplike_proof, atp_proof, outcome),