merged
authorpaulson
Thu, 16 Feb 2023 10:42:39 +0000
changeset 77276 29032b496f2e
parent 77272 0506c3273814 (diff)
parent 77275 386b1b33785c (current diff)
child 77277 c6b50597abbc
merged
--- a/src/HOL/Data_Structures/RBT_Map.thy	Thu Feb 16 10:42:28 2023 +0000
+++ b/src/HOL/Data_Structures/RBT_Map.thy	Thu Feb 16 10:42:39 2023 +0000
@@ -24,11 +24,11 @@
 
 fun del :: "'a::linorder \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt" where
 "del x Leaf = Leaf" |
-"del x (Node l ((a,b), c) r) = (case cmp x a of
+"del x (Node l (ab, _) r) = (case cmp x (fst ab) of
      LT \<Rightarrow> if l \<noteq> Leaf \<and> color l = Black
-           then baldL (del x l) (a,b) r else R (del x l) (a,b) r |
+           then baldL (del x l) ab r else R (del x l) ab r |
      GT \<Rightarrow> if r \<noteq> Leaf\<and> color r = Black
-           then baldR l (a,b) (del x r) else R l (a,b) (del x r) |
+           then baldR l ab (del x r) else R l ab (del x r) |
   EQ \<Rightarrow> join l r)"
 
 definition delete :: "'a::linorder \<Rightarrow> ('a*'b) rbt \<Rightarrow> ('a*'b) rbt" where
@@ -46,10 +46,14 @@
   "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
 by(simp add: update_def inorder_upd inorder_paint)
 
+(* This lemma became necessary below when \<open>del\<close> was converted from pattern-matching to \<open>fst\<close> *)
+lemma del_list_id: "\<forall>ab\<in>set ps. y < fst ab \<Longrightarrow> x \<le> y \<Longrightarrow> del_list x ps = ps"
+by(rule del_list_idem) auto
+
 lemma inorder_del:
  "sorted1(inorder t) \<Longrightarrow>  inorder(del x t) = del_list x (inorder t)"
 by(induction x t rule: del.induct)
-  (auto simp: del_list_simps inorder_join inorder_baldL inorder_baldR)
+  (auto simp: del_list_simps del_list_id inorder_join inorder_baldL inorder_baldR)
 
 lemma inorder_delete:
   "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
@@ -81,19 +85,19 @@
    (color t = Red \<and> bheight (del x t) = bheight t \<and> invc (del x t) \<or>
     color t = Black \<and> bheight (del x t) = bheight t - 1 \<and> invc2 (del x t))"
 proof (induct x t rule: del.induct)
-case (2 x _ y _ c)
-  have "x = y \<or> x < y \<or> x > y" by auto
+case (2 x _ ab c)
+  have "x = fst ab \<or> x < fst ab \<or> x > fst ab" by auto
   thus ?case proof (elim disjE)
-    assume "x = y"
+    assume "x = fst ab"
     with 2 show ?thesis
     by (cases c) (simp_all add: invh_join invc_join)
   next
-    assume "x < y"
+    assume "x < fst ab"
     with 2 show ?thesis
       by(cases c)
         (auto simp: invh_baldL_invc invc_baldL invc2_baldL dest: neq_LeafD)
   next
-    assume "y < x"
+    assume "fst ab < x"
     with 2 show ?thesis
       by(cases c)
         (auto simp: invh_baldR_invc invc_baldR invc2_baldR dest: neq_LeafD)
--- a/src/HOL/Tools/ATP/atp_proof.ML	Thu Feb 16 10:42:28 2023 +0000
+++ b/src/HOL/Tools/ATP/atp_proof.ML	Thu Feb 16 10:42:39 2023 +0000
@@ -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
--- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Thu Feb 16 10:42:28 2023 +0000
+++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Thu Feb 16 10:42:39 2023 +0000
@@ -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) *)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML	Thu Feb 16 10:42:28 2023 +0000
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML	Thu Feb 16 10:42:39 2023 +0000
@@ -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),