improved KBO weights -- beware of explicit applications
authorblanchet
Thu, 09 Feb 2012 14:04:17 +0100
changeset 46443 c86276014571
parent 46442 1e07620d724c
child 46444 db6d2a89a21f
improved KBO weights -- beware of explicit applications
src/HOL/Tools/ATP/atp_problem.ML
src/HOL/Tools/ATP/atp_problem_generate.ML
--- a/src/HOL/Tools/ATP/atp_problem.ML	Thu Feb 09 12:57:59 2012 +0100
+++ b/src/HOL/Tools/ATP/atp_problem.ML	Thu Feb 09 14:04:17 2012 +0100
@@ -754,7 +754,8 @@
 
 fun avoid_clash_with_dfg_keywords s =
   let val n = String.size s in
-    if n < 2 orelse String.isSubstring "_" s then
+    if n < 2 orelse (n = 2 andalso String.sub (s, 0) = String.sub (s, 1)) orelse
+       String.isSubstring "_" s then
       s
     else
       String.substring (s, 0, n - 1) ^
--- a/src/HOL/Tools/ATP/atp_problem_generate.ML	Thu Feb 09 12:57:59 2012 +0100
+++ b/src/HOL/Tools/ATP/atp_problem_generate.ML	Thu Feb 09 14:04:17 2012 +0100
@@ -2706,14 +2706,27 @@
 val use_topo_for_kbo = false (* experimental *)
 val default_kbo_weight = 1
 
+fun have_head_rolling (ATerm (s, args)) =
+    (* ugly hack: may make innocent victims *)
+    if String.isPrefix app_op_name s andalso length args = 2 then
+      have_head_rolling (hd args) ||> append (tl args)
+    else
+      (s, args)
+  | have_head_rolling _ = ("", [])
+
 fun atp_problem_kbo_weights problem =
   let
     fun add_term_deps head (ATerm (s, args)) =
         is_tptp_user_symbol s ? perhaps (try (Graph.add_edge_acyclic (s, head)))
         #> fold (add_term_deps head) args
       | add_term_deps head (AAbs (_, tm)) = add_term_deps head tm
-    fun add_eq_deps (ATerm (s, [ATerm (head, args as _ :: _), rhs])) =
-        is_tptp_equal s ? fold (add_term_deps head) (rhs :: args)
+    fun add_eq_deps (ATerm (s, [lhs as _, rhs])) =
+        if is_tptp_equal s then
+          let val (head, rest) = have_head_rolling lhs in
+            fold (add_term_deps head) (rhs :: rest)
+          end
+        else
+          I
       | add_eq_deps _ = I
     fun add_line_deps _ (Decl (_, s, ty)) =
         is_function_type ty ? Graph.default_node (s, ())