]> matita.cs.unibo.it Git - helm.git/commitdiff
New age selection
authordenes <??>
Thu, 2 Jul 2009 14:20:38 +0000 (14:20 +0000)
committerdenes <??>
Thu, 2 Jul 2009 14:20:38 +0000 (14:20 +0000)
New demodulation strategy

helm/software/components/ng_paramodulation/.depend
helm/software/components/ng_paramodulation/nCicProof.ml
helm/software/components/ng_paramodulation/paramod.ml
helm/software/components/ng_paramodulation/pp.ml
helm/software/components/ng_paramodulation/superposition.ml
helm/software/components/ng_paramodulation/terms.ml
helm/software/components/ng_paramodulation/terms.mli

index ead11bcc5169255404b47d27fa86ac7045cb0199..b78122fff4b3432436603a190ad995f15543b76b 100644 (file)
@@ -11,8 +11,8 @@ cicBlob.cmi: terms.cmi
 nCicProof.cmi: terms.cmi 
 terms.cmo: terms.cmi 
 terms.cmx: terms.cmi 
-pp.cmo: terms.cmi foUtils.cmi pp.cmi 
-pp.cmx: terms.cmx foUtils.cmx pp.cmi 
+pp.cmo: terms.cmi pp.cmi 
+pp.cmx: terms.cmx pp.cmi 
 foSubst.cmo: terms.cmi foSubst.cmi 
 foSubst.cmx: terms.cmx foSubst.cmi 
 orderings.cmo: terms.cmi orderings.cmi 
index 0dbaa1a8f3261d39f601cb4aed85076036d600c7..9223d1bcb12ae0b106bcd5260c5349e0473726f1 100644 (file)
        t vl  
     in
     let get_literal id =
-      let (_, lit, vl, proof),_ = Terms.get_from_bag id bag in
+      let (_, lit, vl, proof),_,_ = Terms.get_from_bag id bag in
       let lit =match lit with 
           | Terms.Predicate t -> assert false 
           | Terms.Equation (l,r,ty,_) -> 
index 0dafc864e215df8985e938b8edf4958de74e29b3..983dd53e84bd28e567904e2c51c74fe8209db52a 100644 (file)
@@ -29,6 +29,7 @@ module Paramod (B : Terms.Blob) = struct
   module IDX = Index.Index(B) 
   module Sup = Superposition.Superposition(B) 
   module Utils = FoUtils.Utils(B) 
+  module Order = Orderings.Orderings(B)
   module WeightOrderedPassives =
       struct
         type t = B.t Terms.passive_clause
@@ -136,19 +137,20 @@ module Paramod (B : Terms.Blob) = struct
   ;;
 
   let backward_infer_step bag maxvar actives passives
-                          g_actives g_passives g_current =
+                          g_actives g_passives g_current iterno =
     (* superposition left, simplifications on goals *)
       debug "infer_left step...";
       let bag, maxvar, new_goals = 
         Sup.infer_left bag maxvar g_current actives 
       in
         debug "Performed infer_left step";
+       let bag = Terms.replace_in_bag (g_current,false,iterno) bag in
         bag, maxvar, actives, passives, g_current::g_actives,
     (add_passive_goals g_passives new_goals)
   ;;
 
   let forward_infer_step bag maxvar actives passives g_actives
-                         g_passives current =
+                         g_passives current iterno =
     (* forward step *)
     
     (* e = select P           *
@@ -172,7 +174,7 @@ module Paramod (B : Terms.Blob) = struct
                Sup.simplify_goal ~no_demod:false maxvar (snd actives) bag acc c
              with
                | None -> bag, acc
-               | Some (bag,c) -> bag,c::acc)
+               | Some (bag,c1) -> bag,if c==c1 then c::acc else c::c1::acc)
           (bag,[]) g_actives 
       in
       let ctable = IDX.index_unit_clause IDX.DT.empty current in
@@ -183,13 +185,14 @@ module Paramod (B : Terms.Blob) = struct
                bag,m,ng@acc) 
           (bag,maxvar,[]) g_actives 
       in
+      let bag = Terms.replace_in_bag (current,false,iterno) bag in
     bag, maxvar, actives,
     add_passive_clauses passives new_clauses, g_actives,
     add_passive_goals g_passives new_goals
   ;;
  
   let rec given_clause ~noinfer 
-    bag maxvar iterno max_steps timeout 
+    bag maxvar iterno weight_picks max_steps timeout 
     actives passives g_actives g_passives 
   =
     let iterno = iterno + 1 in
@@ -219,18 +222,20 @@ module Paramod (B : Terms.Blob) = struct
         else if false then (* activates last chance strategy *)
           begin
            debug("Last chance: "^string_of_float (Unix.gettimeofday()));
-           given_clause ~noinfer:true bag maxvar iterno max_steps 
+           given_clause ~noinfer:true bag maxvar iterno weight_picks max_steps 
              (Some (Unix.gettimeofday () +. 20.))
              actives passives g_actives g_passives;
            raise (Stop (Timeout (maxvar,bag)));
           end
         else raise (Stop (Timeout (maxvar,bag)));
 
-    let use_age = iterno mod 5 = 0 in
+    let use_age = weight_picks = (iterno / 10 + 1) in
+    let weight_picks = if use_age then 0 else weight_picks+1
+    in
 
     let rec aux_select bag passives g_passives =
       let backward,(weight,current),passives,g_passives =
-        select ~use_age:false passives g_passives
+        select ~use_age passives g_passives
       in
        if use_age && weight > monster then
          let bag,cl = Terms.add_to_bag current bag in
@@ -239,6 +244,7 @@ module Paramod (B : Terms.Blob) = struct
            else
              aux_select bag (add_passive_clause passives cl) g_passives
        else
+         let bag = Terms.replace_in_bag (current,false,iterno) bag in
         if backward then
           let _ = debug ("Selected goal : " ^ Pp.pp_unit_clause current) in
          match 
@@ -255,7 +261,7 @@ module Paramod (B : Terms.Blob) = struct
                  bag,maxvar,actives,passives,g_actives,g_passives
                else
                  backward_infer_step bag maxvar actives passives
-                   g_actives g_passives g_current
+                   g_actives g_passives g_current iterno
         else
           let _ = debug ("Selected fact : " ^ Pp.pp_unit_clause current) in
          (*let is_orphan = Sup.orphan_murder bag (fst actives) current in*)
@@ -265,7 +271,7 @@ module Paramod (B : Terms.Blob) = struct
               else  bag, Some (current,actives)
             else if Sup.orphan_murder bag (fst actives) current then
              let _ = debug "Orphan murdered" in
-              let bag = Terms.replace_in_bag (current,true) bag in
+              let bag = Terms.replace_in_bag (current,true,iterno) bag in
                 bag, None
             else Sup.keep_simplified current actives bag maxvar
           with
@@ -291,7 +297,7 @@ module Paramod (B : Terms.Blob) = struct
                     bag,maxvar,actives,passives,g_actives,g_passives
                   else
                     forward_infer_step bag maxvar actives passives
-                      g_actives g_passives current
+                      g_actives g_passives current iterno
     in
     
 
@@ -314,7 +320,7 @@ module Paramod (B : Terms.Blob) = struct
         (Printf.sprintf "Number of passives : %d"
            (passive_set_cardinal passives));
       given_clause ~noinfer
-        bag maxvar iterno max_steps timeout 
+        bag maxvar iterno weight_picks max_steps timeout 
         actives passives g_actives g_passives
   ;;
 
@@ -330,16 +336,16 @@ module Paramod (B : Terms.Blob) = struct
     let actives = [], IDX.DT.empty in
     try 
      given_clause ~noinfer:false
-      bag maxvar 0 max_steps timeout actives passives g_actives g_passives
+      bag maxvar 0  0 max_steps timeout actives passives g_actives g_passives
     with 
     | Sup.Success (bag, _, (i,_,_,_)) ->
         let l =
           let rec traverse ongoal (accg,acce) i =
             match Terms.get_from_bag i bag with
-              | (id,_,_,Terms.Exact _),_ ->
+              | (id,_,_,Terms.Exact _),_,_ ->
                   if ongoal then [i],acce else
                     if (List.mem i acce) then accg,acce else accg,acce@[i]
-              | (_,_,_,Terms.Step (_,i1,i2,_,_,_)),_ ->
+              | (_,_,_,Terms.Step (_,i1,i2,_,_,_)),_,_ ->
                   if (not ongoal) && (List.mem i acce) then accg,acce
                   else
                     let accg,acce = 
@@ -350,9 +356,23 @@ module Paramod (B : Terms.Blob) = struct
           let gsteps,esteps = traverse true ([],[]) i in
             (List.rev esteps)@gsteps
         in
-(*       List.iter (fun id -> let (cl,d) =
-                      Terms.M.find id bag in 
-                   if d then prerr_endline (Pp.pp_unit_clause cl)) l;*)
+       let max_w = List.fold_left (fun acc i ->
+                   let (cl,_,_) = Terms.get_from_bag i bag in
+                   max acc (Order.compute_unit_clause_weight cl)) 0 l in
+         prerr_endline "Statistics :";
+         prerr_endline ("Max weight : " ^ (string_of_int max_w));
+(*       List.iter (fun id -> let ((_,lit,_,proof as cl),d,it) =
+           Terms.get_from_bag id bag in
+             if d then
+               prerr_endline
+               (Printf.sprintf "Id : %d, selected at %d, weight %d,disc, by %s"
+                  id it (Order.compute_unit_clause_weight cl) 
+                  (Pp.pp_proof_step proof))
+             else
+              prerr_endline
+               (Printf.sprintf "Id : %d, selected at %d, weight %d by %s"
+                  id it (Order.compute_unit_clause_weight cl) 
+                  (Pp.pp_proof_step proof))) l;*)
         prerr_endline 
           (Printf.sprintf "Found proof, %fs" 
             (Unix.gettimeofday() -. initial_timestamp));
index 1c9fb2fc636018ac07d8d6e450f1de9b971378a1..625de5c31994b08f8b6fc0c08a85c5b1c620584a 100644 (file)
@@ -68,8 +68,8 @@ let pp_proof bag ~formatter:f p =
           eq (string_of_rule rule);
        Format.fprintf f "|%d with %d dir %s))" eq1 eq2
          (string_of_direction dir);
-       let (_, _, _, proof1),_ = Terms.get_from_bag eq1 bag in
-       let (_, _, _, proof2),_ = Terms.get_from_bag eq2 bag in
+       let (_, _, _, proof1),_,_ = Terms.get_from_bag eq1 bag in
+       let (_, _, _, proof2),_,_ = Terms.get_from_bag eq2 bag in
          Format.fprintf f "@[<v 2>";
           aux eq1 proof1;          
           aux eq2 proof2;
@@ -123,7 +123,7 @@ let pp_unit_clause ~formatter:f c =
 let pp_bag ~formatter:f (_,bag) = 
   Format.fprintf f "@[<v>";
   Terms.M.iter 
-  (fun _ (c,d) -> pp_unit_clause ~formatter:f c;
+  (fun _ (c,d,_) -> pp_unit_clause ~formatter:f c;
      if d then Format.fprintf f " (discarded)@;"
      else Format.fprintf f "@;") bag;
   Format.fprintf f "@]"
index 96b16bb07b6a66dd674495d750e76075185faae2..c1d5f8bdcffcd9ab3d7a2f226a278e0afe4c6440 100644 (file)
@@ -31,12 +31,16 @@ module Superposition (B : Terms.Blob) =
     ;;
 
     let first_position pos ctx t f =
+      let inject_pos pos ctx = function
+       | None -> None
+       | Some (a,b,c,d,e) -> Some(ctx a,b,c,d,e,pos)
+      in
       let rec aux pos ctx = function
-      | Terms.Leaf _ as t -> f t pos ctx 
+      | Terms.Leaf _ as t -> inject_pos pos ctx (f t)
       | Terms.Var _ -> None
       | Terms.Node l as t->
-          match f t pos ctx with
-          | Some _ as x -> x
+          match f t with
+          | Some _ as x -> inject_pos pos ctx x
           | None ->
               let rec first pre post = function
                 | [] -> None
@@ -72,6 +76,29 @@ module Superposition (B : Terms.Blob) =
         aux pos ctx t
     ;;
 
+    let parallel_positions bag pos ctx id t f =
+      let rec aux bag pos ctx id = function
+      | Terms.Leaf _ as t -> f bag t pos ctx id
+      | Terms.Var _ as t -> bag,t,id
+      | Terms.Node l as t->
+          let bag,t,id1 = f bag t pos ctx id in
+           if id = id1 then
+              let bag, l, _, id = 
+               List.fold_left
+                 (fun (bag,pre,post,id) t ->
+                     let newctx = fun x -> ctx (Terms.Node (pre@[x]@post)) in
+                    let newpos = (List.length pre)::pos in
+                    let bag,newt,id = aux bag newpos newctx id t in
+                      if post = [] then bag, pre@[newt], [], id
+                       else bag, pre @ [newt], List.tl post, id)
+                 (bag, [], List.tl l, id) l
+              in
+               bag, Terms.Node l, id
+           else bag,t,id1
+      in
+        aux bag pos ctx id t
+    ;;
+
     let vars_of_term t =
       let rec aux acc = function
        | Terms.Leaf _ -> acc
@@ -102,7 +129,7 @@ module Superposition (B : Terms.Blob) =
     
     (* ============ simplification ================= *)
 
-    let demod table varlist subterm pos context =
+    let demod table varlist subterm =
       let cands = IDX.DT.retrieve_generalizations table subterm in
       list_first
         (fun (dir, (id,lit,vl,_)) ->
@@ -114,25 +141,25 @@ module Superposition (B : Terms.Blob) =
                  let subst, varlist = 
                    Unif.unification (varlist@vl) varlist subterm side 
                  in
+                 let side = Subst.apply_subst subst side in
+                 let newside = Subst.apply_subst subst newside in
                  if o = Terms.Incomparable then
-                   let side = Subst.apply_subst subst side in
-                   let newside = Subst.apply_subst subst newside in
                    let o = Order.compare_terms newside side in
                    (* Riazanov, pp. 45 (ii) *)
                    if o = Terms.Lt then
-                     Some (context newside, subst, varlist, id, pos, dir)
+                     Some (newside, subst, varlist, id, dir)
                    else 
                      ((*prerr_endline ("Filtering: " ^ 
                         Pp.pp_foterm side ^ " =(< || =)" ^ 
                         Pp.pp_foterm newside ^ " coming from " ^ 
                         Pp.pp_unit_clause uc );*)None)
                  else
-                   Some (context newside, subst, varlist, id, pos, dir)
+                   Some (newside, subst, varlist, id, dir)
                with FoUnif.UnificationFailure _ -> None)
         (IDX.ClauseSet.elements cands)
     ;;
 
-    let demodulate_once ~jump_to_right bag (id, literal, vl, pr) table =
+    let demodulate_once_old ~jump_to_right bag (id, literal, vl, pr) table =
       match literal with
       | Terms.Predicate t -> assert false
       | Terms.Equation (l,r,ty,_) ->
@@ -142,7 +169,7 @@ module Superposition (B : Terms.Blob) =
            (demod table vl)
        in
         match left_position with
-         | Some (newt, subst, varlist, id2, pos, dir) ->
+         | Some (newt, subst, varlist, id2, dir, pos) ->
              begin
                match build_clause bag (fun _ -> true) Terms.Demodulation 
                  newt subst varlist id id2 pos dir
@@ -156,7 +183,7 @@ module Superposition (B : Terms.Blob) =
                (demod table vl)
              with
                | None -> None
-               | Some (newt, subst, varlist, id2, pos, dir) ->
+               | Some (newt, subst, varlist, id2, dir, pos) ->
                    match build_clause bag (fun _ -> true)
                      Terms.Demodulation newt subst varlist id id2 pos dir
                    with
@@ -164,6 +191,39 @@ module Superposition (B : Terms.Blob) =
                        | Some x -> Some (x,true)
     ;;
 
+    let parallel_demod table vl bag t pos ctx id =
+      match demod table vl t with
+       | None -> (bag,t,id)
+       | Some (newside, subst, vl, id2, dir) ->
+           match build_clause bag (fun _ -> true)
+             Terms.Demodulation (ctx newside) subst vl id id2 pos dir
+           with
+             | None -> assert false
+             | Some (bag,(id,_,_,_)) ->
+                   (bag,newside,id)
+    ;;
+
+    let demodulate_once ~jump_to_right bag (id, literal, vl, pr) table =
+      match literal with
+      | Terms.Predicate t -> assert false
+      | Terms.Equation (l,r,ty,_) ->
+         let bag,l,id1 = if jump_to_right then (bag,l,id) else
+           parallel_positions bag [2]
+             (fun x -> Terms.Node [ Terms.Leaf B.eqP; ty; x; r ]) id l
+             (parallel_demod table vl)
+         in
+         let jump_to_right = id1 = id in
+         let bag,r,id2 =
+           parallel_positions bag [3]
+             (fun x -> Terms.Node [ Terms.Leaf B.eqP; ty; l; x ]) id1 r
+             (parallel_demod table vl)
+         in
+           if id = id2 then None
+           else
+             let cl,_,_ = Terms.get_from_bag id2 bag in
+               Some ((bag,cl),jump_to_right)
+    ;;
+
     let rec demodulate ~jump_to_right bag clause table =
       match demodulate_once ~jump_to_right bag clause table with
       | None -> bag, clause
@@ -171,8 +231,37 @@ module Superposition (B : Terms.Blob) =
          bag clause table
     ;;
 
-    let demodulate bag clause table = demodulate ~jump_to_right:false
-      bag clause table
+    let rec demodulate_old ~jump_to_right bag clause table =
+      match demodulate_once_old ~jump_to_right bag clause table with
+       | None -> bag, clause
+       | Some ((bag, clause),r) -> demodulate_old ~jump_to_right:r
+         bag clause table
+    ;;
+
+    let are_alpha_eq cl1 cl2 =
+      let get_term (_,lit,_,_) =
+       match lit with
+         | Terms.Predicate _ -> assert false
+         | Terms.Equation (l,r,ty,_) ->
+             Terms.Node [Terms.Leaf B.eqP; ty; l ; r]
+      in
+       try ignore(Unif.alpha_eq (get_term cl1) (get_term cl2)) ; true
+       with FoUnif.UnificationFailure _ -> false
+    ;;
+
+    let demodulate bag clause table =
+(*      let (bag1,c1), (_,c2) =*)
+       demodulate ~jump_to_right:false bag clause table
+(*     demodulate_old ~jump_to_right:false bag clause table*)
+(*      in
+       if are_alpha_eq c1 c2 then bag1,c1
+       else begin
+         prerr_endline (Pp.pp_unit_clause c1);
+         prerr_endline (Pp.pp_unit_clause c2);
+         prerr_endline "Bag :";
+         prerr_endline (Pp.pp_bag bag1);
+         assert false
+       end*)
     ;;
 
     (* move away *)
@@ -303,9 +392,9 @@ module Superposition (B : Terms.Blob) =
 
     let rec orphan_murder bag acc i =
       match Terms.get_from_bag i bag with
-       | (_,_,_,Terms.Exact _),discarded -> (discarded,acc)
-       | (_,_,_,Terms.Step (_,i1,i2,_,_,_)),true -> (true,acc)
-       | (_,_,_,Terms.Step (_,i1,i2,_,_,_)),false ->
+       | (_,_,_,Terms.Exact _),discarded,_ -> (discarded,acc)
+       | (_,_,_,Terms.Step (_,i1,i2,_,_,_)),true,_ -> (true,acc)
+       | (_,_,_,Terms.Step (_,i1,i2,_,_,_)),false,_ ->
             if (List.mem i acc) then (false,acc)
             else match orphan_murder bag acc i1 with
              | (true,acc) -> (true,acc)
@@ -336,7 +425,9 @@ module Superposition (B : Terms.Blob) =
     let simplify table maxvar bag clause =
       match simplify table maxvar bag clause with
        | bag, None ->
-           Terms.replace_in_bag (clause,true) bag, None
+           let (id,_,_,_) = clause in
+           let (_,_,iter) = Terms.get_from_bag id bag in
+           Terms.replace_in_bag (clause,true,iter) bag, None
        | bag, Some clause -> bag, Some clause
     (*let (id,_,_,_) = clause in
            if orphan_murder bag clause then
@@ -432,17 +523,6 @@ module Superposition (B : Terms.Blob) =
        keep_simplified_aux ~new_cl:true cl (alist,atable) bag []
     ;;
 
-    let are_alpha_eq cl1 cl2 =
-      let get_term (_,lit,_,_) =
-       match lit with
-         | Terms.Predicate _ -> assert false
-         | Terms.Equation (l,r,ty,_) ->
-             Terms.Node [Terms.Leaf B.eqP; ty; l ; r]
-      in
-       try ignore(Unif.alpha_eq (get_term cl1) (get_term cl2)) ; true
-       with FoUnif.UnificationFailure _ -> false
-;;
-
     (* this is like simplify but raises Success *)
     let simplify_goal ~no_demod maxvar table bag g_actives clause = 
       let bag, clause = 
@@ -500,8 +580,7 @@ module Superposition (B : Terms.Blob) =
                    else 
                      ((*prerr_endline ("Filtering: " ^ 
                         Pp.pp_foterm side ^ " =(< || =)" ^ 
-                        Pp.pp_foterm newside ^ " coming from " ^ 
-                        Pp.pp_unit_clause uc );*)None)
+                        Pp.pp_foterm newside);*)None)
                  else
                    Some (context newside, subst, varlist, id, pos, dir)
                with FoUnif.UnificationFailure _ -> None)
index aefbc01a864ea3a870746f11d7273f27d5eef8ff..169d19b80b68a52d0eb062162464e451a96a5757 100644 (file)
@@ -57,16 +57,16 @@ module M : Map.S with type key = int
   = Map.Make(OT) 
 
 type 'a bag = int
-              * (('a unit_clause * bool) M.t)
+              * (('a unit_clause * bool * int) M.t)
 
   let add_to_bag (_,lit,vl,proof) (id,bag) =
     let id = id+1 in
     let clause = (id, lit, vl, proof) in
-    let bag = M.add id (clause,false) bag in
+    let bag = M.add id (clause,false,0) bag in
     (id,bag), clause 
    ;;
 
-  let replace_in_bag ((id,_,_,_),_ as cl) (max_id,bag) =
+  let replace_in_bag ((id,_,_,_),_,_ as cl) (max_id,bag) =
     let bag = M.add id cl bag in
       (max_id,bag)
   ;;
index 89d6993d908ada394458ce40d97f86a0f061bc60..fe03cc454fc88120faf22837305477c1d342d7f5 100644 (file)
@@ -55,7 +55,7 @@ type 'a passive_clause = int * 'a unit_clause (* weight * equation *)
 module M : Map.S with type key = int 
 
 type 'a bag = int (* max ID  *)
-              * (('a unit_clause * bool) M.t)
+              * (('a unit_clause * bool * int) M.t)
 
 (* also gives a fresh ID to the clause *)
     val add_to_bag : 
@@ -63,11 +63,11 @@ type 'a bag = int (* max ID  *)
             'a bag * 'a unit_clause
 
     val replace_in_bag : 
-          'a unit_clause * bool -> 'a bag ->
+          'a unit_clause * bool * int -> 'a bag ->
             'a bag
 
     val get_from_bag : 
-          int -> 'a bag -> 'a unit_clause * bool
+          int -> 'a bag -> 'a unit_clause * bool * int
 
     val empty_bag : 'a bag