]> matita.cs.unibo.it Git - helm.git/commitdiff
code refactoring for paramodulation
authorEnrico Tassi <enrico.tassi@inria.fr>
Thu, 25 Jun 2009 10:17:36 +0000 (10:17 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Thu, 25 Jun 2009 10:17:36 +0000 (10:17 +0000)
14 files changed:
helm/software/components/ng_paramodulation/.depend
helm/software/components/ng_paramodulation/.depend.opt
helm/software/components/ng_paramodulation/Makefile
helm/software/components/ng_paramodulation/cicBlob.ml
helm/software/components/ng_paramodulation/nCicBlob.ml
helm/software/components/ng_paramodulation/nCicParamod.ml [new file with mode: 0644]
helm/software/components/ng_paramodulation/nCicParamod.mli [new file with mode: 0644]
helm/software/components/ng_paramodulation/nCicProof.ml [new file with mode: 0644]
helm/software/components/ng_paramodulation/nCicProof.mli [new file with mode: 0644]
helm/software/components/ng_paramodulation/paramod.ml
helm/software/components/ng_paramodulation/paramod.mli
helm/software/components/ng_paramodulation/terms.ml
helm/software/components/ng_paramodulation/terms.mli
helm/software/components/ng_tactics/nTactics.ml

index 54d1bb2667f92c350cd5127c983e80a95bfb7f3a..a680eb8efa1cddba54b094775dc267137b049d57 100644 (file)
@@ -1,3 +1,4 @@
+terms.cmi: 
 pp.cmi: terms.cmi 
 foSubst.cmi: terms.cmi 
 orderings.cmi: terms.cmi 
@@ -5,8 +6,11 @@ foUtils.cmi: terms.cmi
 index.cmi: terms.cmi 
 foUnif.cmi: terms.cmi 
 superposition.cmi: terms.cmi index.cmi 
+paramod.cmi: terms.cmi 
 nCicBlob.cmi: terms.cmi 
 cicBlob.cmi: terms.cmi 
+nCicProof.cmi: terms.cmi 
+nCicParamod.cmi: 
 terms.cmo: terms.cmi 
 terms.cmx: terms.cmi 
 pp.cmo: terms.cmi pp.cmi 
@@ -25,11 +29,17 @@ superposition.cmo: terms.cmi pp.cmi orderings.cmi index.cmi foUtils.cmi \
     foUnif.cmi foSubst.cmi superposition.cmi 
 superposition.cmx: terms.cmx pp.cmx orderings.cmx index.cmx foUtils.cmx \
     foUnif.cmx foSubst.cmx superposition.cmi 
-nCicBlob.cmo: terms.cmi foUtils.cmi foSubst.cmi nCicBlob.cmi 
-nCicBlob.cmx: terms.cmx foUtils.cmx foSubst.cmx nCicBlob.cmi 
+paramod.cmo: terms.cmi superposition.cmi pp.cmi index.cmi foUtils.cmi \
+    foUnif.cmi paramod.cmi 
+paramod.cmx: terms.cmx superposition.cmx pp.cmx index.cmx foUtils.cmx \
+    foUnif.cmx paramod.cmi 
+nCicBlob.cmo: terms.cmi foUtils.cmi nCicBlob.cmi 
+nCicBlob.cmx: terms.cmx foUtils.cmx nCicBlob.cmi 
 cicBlob.cmo: terms.cmi cicBlob.cmi 
 cicBlob.cmx: terms.cmx cicBlob.cmi 
-paramod.cmo: terms.cmi superposition.cmi pp.cmi nCicBlob.cmi index.cmi \
-    foUtils.cmi foUnif.cmi paramod.cmi 
-paramod.cmx: terms.cmx superposition.cmx pp.cmx nCicBlob.cmx index.cmx \
-    foUtils.cmx foUnif.cmx paramod.cmi 
+nCicProof.cmo: terms.cmi foSubst.cmi nCicProof.cmi 
+nCicProof.cmx: terms.cmx foSubst.cmx nCicProof.cmi 
+nCicParamod.cmo: terms.cmi paramod.cmi nCicProof.cmi nCicBlob.cmi \
+    nCicParamod.cmi 
+nCicParamod.cmx: terms.cmx paramod.cmx nCicProof.cmx nCicBlob.cmx \
+    nCicParamod.cmi 
index 666cba23e955cea150f8f54ef68a9ad1f6941244..a680eb8efa1cddba54b094775dc267137b049d57 100644 (file)
@@ -1,3 +1,4 @@
+terms.cmi: 
 pp.cmi: terms.cmi 
 foSubst.cmi: terms.cmi 
 orderings.cmi: terms.cmi 
@@ -5,8 +6,11 @@ foUtils.cmi: terms.cmi
 index.cmi: terms.cmi 
 foUnif.cmi: terms.cmi 
 superposition.cmi: terms.cmi index.cmi 
+paramod.cmi: terms.cmi 
 nCicBlob.cmi: terms.cmi 
 cicBlob.cmi: terms.cmi 
+nCicProof.cmi: terms.cmi 
+nCicParamod.cmi: 
 terms.cmo: terms.cmi 
 terms.cmx: terms.cmi 
 pp.cmo: terms.cmi pp.cmi 
@@ -25,11 +29,17 @@ superposition.cmo: terms.cmi pp.cmi orderings.cmi index.cmi foUtils.cmi \
     foUnif.cmi foSubst.cmi superposition.cmi 
 superposition.cmx: terms.cmx pp.cmx orderings.cmx index.cmx foUtils.cmx \
     foUnif.cmx foSubst.cmx superposition.cmi 
+paramod.cmo: terms.cmi superposition.cmi pp.cmi index.cmi foUtils.cmi \
+    foUnif.cmi paramod.cmi 
+paramod.cmx: terms.cmx superposition.cmx pp.cmx index.cmx foUtils.cmx \
+    foUnif.cmx paramod.cmi 
 nCicBlob.cmo: terms.cmi foUtils.cmi nCicBlob.cmi 
 nCicBlob.cmx: terms.cmx foUtils.cmx nCicBlob.cmi 
 cicBlob.cmo: terms.cmi cicBlob.cmi 
 cicBlob.cmx: terms.cmx cicBlob.cmi 
-paramod.cmo: terms.cmi superposition.cmi pp.cmi nCicBlob.cmi index.cmi \
-    foUtils.cmi foUnif.cmi paramod.cmi 
-paramod.cmx: terms.cmx superposition.cmx pp.cmx nCicBlob.cmx index.cmx \
-    foUtils.cmx foUnif.cmx paramod.cmi 
+nCicProof.cmo: terms.cmi foSubst.cmi nCicProof.cmi 
+nCicProof.cmx: terms.cmx foSubst.cmx nCicProof.cmi 
+nCicParamod.cmo: terms.cmi paramod.cmi nCicProof.cmi nCicBlob.cmi \
+    nCicParamod.cmi 
+nCicParamod.cmx: terms.cmx paramod.cmx nCicProof.cmx nCicBlob.cmx \
+    nCicParamod.cmi 
index 6fff2162c7f60ea9cfbab8311d3eabea75dd042c..a8d0b8ea9f07ae53c0ea99da97f24661fcd5e97d 100644 (file)
@@ -3,7 +3,7 @@ PACKAGE = ng_paramodulation
 INTERFACE_FILES = \
        terms.mli pp.mli foSubst.mli \
        orderings.mli foUtils.mli index.mli foUnif.mli superposition.mli \
-       nCicBlob.mli cicBlob.mli paramod.mli
+       paramod.mli nCicBlob.mli cicBlob.mli nCicProof.mli nCicParamod.mli
 
 IMPLEMENTATION_FILES = $(INTERFACE_FILES:%.mli=%.ml)
 
index 888e6c0013aa30962682e38fd4d35aeac71de9cb..5c672cddfd0e9614dd905242b991a5aa7e5da6fe 100644 (file)
@@ -38,6 +38,5 @@ module CicBlob(C : CicContext) : Terms.Blob with type t = Cic.term = struct
 
   let saturate = assert false;;
 
-  let mk_proof = assert false;;
 end
 
index e611266d92cfad1c330a13c98b7834d084e95f94..aa2f4f9397cc31ee4a4066fe2011e526b84157b7 100644 (file)
@@ -84,192 +84,4 @@ module NCicBlob(C : NCicContext) : Terms.Blob with type t = NCic.term = struct
     NCic.Const r
   ;;
 
-  let eq_ind = 
-    let r = 
-      OCic2NCic.reference_of_oxuri 
-       (UriManager.uri_of_string 
-         "cic:/matita/logic/equality/eq_ind.con")
-    in
-    NCic.Const r
-  ;;
-
-  let eq_ind_r = 
-    let r = 
-      OCic2NCic.reference_of_oxuri 
-       (UriManager.uri_of_string 
-         "cic:/matita/logic/equality/eq_elim_r.con")
-    in
-    NCic.Const r
-  ;;
-
-  let eq_refl = 
-    let r = 
-      OCic2NCic.reference_of_oxuri 
-       (UriManager.uri_of_string 
-         "cic:/matita/logic/equality/eq.ind#xpointer(1/1/1)")
-    in
-    NCic.Const r
-  ;;
-
-  let extract lift vl t =
-    let rec pos i = function 
-      | [] -> raise Not_found 
-      | j :: tl when j <> i -> 1+ pos i tl
-      | _ -> 1
-    in
-    let vl_len = List.length vl in
-    let rec extract = function
-      | Terms.Leaf x -> NCicSubstitution.lift (vl_len+lift) x
-      | Terms.Var j -> 
-           (try NCic.Rel (pos j vl) with Not_found -> NCic.Implicit `Term) 
-      | Terms.Node l -> NCic.Appl (List.map extract l)
-    in
-      extract t
-  ;;
-
-   let rec ppfot = function 
-     | Terms.Leaf _ -> "."
-     | Terms.Var i -> "?" ^ string_of_int i
-     | Terms.Node l -> "(" ^ String.concat " " (List.map ppfot l) ^ ")"
-   ;;
-
-   let mk_predicate hole_type amount ft p1 vl =
-    let rec aux t p = 
-      match p with
-      | [] -> NCic.Rel 1
-      | n::tl ->
-          match t with
-          | Terms.Leaf _ 
-          | Terms.Var _ -> 
-               prerr_endline ("term: " ^ ppfot ft);            
-               prerr_endline ("path: " ^ String.concat "," 
-                 (List.map string_of_int p1));
-               assert false
-          | Terms.Node l -> 
-              let l = 
-                HExtlib.list_mapi
-                  (fun t i -> 
-                    if i = n then aux t tl 
-                    else extract amount (0::vl) t)
-                  l
-              in            
-              NCic.Appl l
-    in
-      NCic.Lambda("x", hole_type, aux ft (List.rev p1))
-    ;;
-
-  let mk_proof (bag : NCic.term Terms.bag) mp steps =
-    let module Subst = FoSubst in
-    let position i l = 
-      let rec aux = function
-       | [] -> assert false
-       | (j,_) :: tl when i = j -> 1
-       | _ :: tl -> 1 + aux tl
-      in
-        aux l
-    in
-    let vars_of i l = fst (List.assoc i l) in
-    let ty_of i l = snd (List.assoc i l) in
-    let close_with_lambdas vl t = 
-      List.fold_left 
-       (fun t i -> 
-          NCic.Lambda ("x"^string_of_int i, NCic.Implicit `Type, t))
-       t vl  
-    in
-    let close_with_forall vl t = 
-      List.fold_left 
-       (fun t i -> 
-          NCic.Prod ("x"^string_of_int i, NCic.Implicit `Type, t))
-       t vl  
-    in
-    let get_literal id =
-      let _, lit, vl, proof = Terms.M.find id bag in
-      let lit =match lit with 
-          | Terms.Predicate t -> assert false 
-          | Terms.Equation (l,r,ty,_) -> 
-              Terms.Node [ Terms.Leaf eqP; ty; l; r]
-      in
-       lit, vl, proof
-    in
-    let rec aux ongoal seen = function
-      | [] -> NCic.Rel 1
-      | id :: tl ->
-          let amount = List.length seen in
-          let lit,vl,proof = get_literal id in
-          if not ongoal && id = mp then
-            ((*prerr_endline ("Reached m point, id=" ^ (string_of_int id));*)
-             NCic.LetIn ("clause_" ^ string_of_int id, 
-                extract amount [] lit, 
-                (NCic.Appl [eq_refl;NCic.Implicit `Type;NCic.Implicit `Term]),
-                aux true ((id,([],lit))::seen) (id::tl)))
-          else
-          match proof with
-          | Terms.Exact _ when tl=[] ->
-             (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id));*)
-             aux ongoal seen tl
-          | Terms.Step _ when tl=[] -> assert false
-          | Terms.Exact ft ->
-             (* prerr_endline ("Exact for " ^ (string_of_int id));*)
-               NCic.LetIn ("clause_" ^ string_of_int id, 
-                 close_with_forall vl (extract amount vl lit),
-                 close_with_lambdas vl (extract amount vl ft),
-                 aux ongoal 
-                   ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
-          | Terms.Step (_, id1, id2, dir, pos, subst) ->
-              let id, id1,(lit,vl,proof) =
-               if ongoal then id1,id,get_literal id1
-               else id,id1,(lit,vl,proof)
-             in
-             let vl = if ongoal then Subst.filter subst vl else vl in
-              let proof_of_id id = 
-                let vars = List.rev (vars_of id seen) in
-                let args = List.map (Subst.apply_subst subst) vars in
-                let args = List.map (extract amount vl) args in
-               let rel_for_id = NCic.Rel (List.length vl + position id seen) in
-                 if args = [] then rel_for_id              
-                  else NCic.Appl (rel_for_id::args)
-              in
-              let p_id1 = proof_of_id id1 in
-              let p_id2 = proof_of_id id2 in
-              let pred, hole_type, l, r = 
-                let id1_ty = ty_of id1 seen in
-                let id2_ty,l,r = 
-                  match ty_of id2 seen with
-                  | Terms.Node [ _; t; l; r ] -> 
-                      extract amount vl (Subst.apply_subst subst t),
-                      extract amount vl (Subst.apply_subst subst l),
-                      extract amount vl (Subst.apply_subst subst r)
-                  | _ -> assert false
-                in
-                 (*prerr_endline "mk_predicate :";
-                 if ongoal then prerr_endline "ongoal=true"
-                 else prerr_endline "ongoal=false";
-                 prerr_endline ("id=" ^ string_of_int id);
-                 prerr_endline ("id1=" ^ string_of_int id1);
-                 prerr_endline ("id2=" ^ string_of_int id2);
-                 prerr_endline ("Positions :" ^
-                                  (String.concat ", "
-                                     (List.map string_of_int pos)));*)
-                mk_predicate 
-                  id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
-                id2_ty,
-                l,r
-              in
-              let l, r, eq_ind = 
-                if (ongoal=true) = (dir=Terms.Left2Right) then
-                  r,l,eq_ind_r
-                else
-                  l,r,eq_ind
-              in
-               NCic.LetIn ("clause_" ^ string_of_int id, 
-                 close_with_forall vl (extract amount vl lit),
-                          (* NCic.Implicit `Type, *)
-                 close_with_lambdas vl 
-                   (NCic.Appl [ eq_ind ; hole_type; l; pred; p_id1; r; p_id2 ]),
-                 aux ongoal 
-                   ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
-    in 
-      aux false [] steps
-  ;;
-
  end
diff --git a/helm/software/components/ng_paramodulation/nCicParamod.ml b/helm/software/components/ng_paramodulation/nCicParamod.ml
new file mode 100644 (file)
index 0000000..5fe0a34
--- /dev/null
@@ -0,0 +1,60 @@
+(*
+    ||M||  This file is part of HELM, an Hypertextual, Electronic        
+    ||A||  Library of Mathematics, developed at the Computer Science     
+    ||T||  Department, University of Bologna, Italy.                     
+    ||I||                                                                
+    ||T||  HELM is free software; you can redistribute it and/or         
+    ||A||  modify it under the terms of the GNU General Public License   
+    \   /  version 2 or (at your option) any later version.      
+     \ /   This software is distributed as is, NO WARRANTY.     
+      V_______________________________________________________________ *)
+
+(* $Id: orderings.ml 9869 2009-06-11 22:52:38Z denes $ *)
+
+let nparamod rdb metasenv subst context t table =
+  let module C = 
+    struct
+      let metasenv = metasenv
+      let subst = subst
+      let context = context
+    end
+  in
+  let module B = NCicBlob.NCicBlob(C) in
+  let module P = Paramod.Paramod(B) in
+  let bag, maxvar = Terms.M.empty, 0 in
+  let (bag,maxvar), passives = 
+    HExtlib.list_mapi_acc (fun x _ a -> P.mk_passive a x) (bag,maxvar) table
+  in
+  let (bag,maxvar), goals = 
+    HExtlib.list_mapi_acc (fun x _ a -> P.mk_goal a x) (bag,maxvar) [t]
+  in
+  let solutions = P.paramod (bag,maxvar) ~g_passives:goals ~passives in
+  List.map 
+    (fun (bag,i,l) ->
+      let stamp = Unix.gettimeofday () in
+      let proofterm = NCicProof.mk_proof bag i l in
+      prerr_endline (Printf.sprintf "Got proof term in %fs"
+        (Unix.gettimeofday() -. stamp));
+      let metasenv, proofterm = 
+        let rec aux k metasenv = function
+          | NCic.Meta _ as t -> metasenv, t
+          | NCic.Implicit _ -> 
+              let metasenv,i,_,_=NCicMetaSubst.mk_meta metasenv context `Term in
+              metasenv, NCic.Meta (i,(k,NCic.Irl (List.length context)))
+          | t -> NCicUntrusted.map_term_fold_a 
+                  (fun _ k -> k+1) k aux metasenv t
+        in
+         aux 0 metasenv proofterm
+      in
+      let metasenv, subst, proofterm, _prooftype = 
+        NCicRefiner.typeof 
+          (rdb#set_coerc_db NCicCoercion.empty_db) 
+          metasenv subst context proofterm None
+      in
+      proofterm, metasenv, subst)
+    solutions
+;;
+  
+  
+
+
diff --git a/helm/software/components/ng_paramodulation/nCicParamod.mli b/helm/software/components/ng_paramodulation/nCicParamod.mli
new file mode 100644 (file)
index 0000000..d24e5b5
--- /dev/null
@@ -0,0 +1,18 @@
+(*
+    ||M||  This file is part of HELM, an Hypertextual, Electronic        
+    ||A||  Library of Mathematics, developed at the Computer Science     
+    ||T||  Department, University of Bologna, Italy.                     
+    ||I||                                                                
+    ||T||  HELM is free software; you can redistribute it and/or         
+    ||A||  modify it under the terms of the GNU General Public License   
+    \   /  version 2 or (at your option) any later version.      
+     \ /   This software is distributed as is, NO WARRANTY.     
+      V_______________________________________________________________ *)
+
+(* $Id: orderings.ml 9869 2009-06-11 22:52:38Z denes $ *)
+
+val nparamod :
+  #NRstatus.status ->
+  NCic.metasenv -> NCic.substitution -> NCic.context -> 
+    (NCic.term * NCic.term) -> (NCic.term * NCic.term) list ->
+     (NCic.term * NCic.metasenv * NCic.substitution) list
diff --git a/helm/software/components/ng_paramodulation/nCicProof.ml b/helm/software/components/ng_paramodulation/nCicProof.ml
new file mode 100644 (file)
index 0000000..35c3c02
--- /dev/null
@@ -0,0 +1,204 @@
+(*
+    ||M||  This file is part of HELM, an Hypertextual, Electronic        
+    ||A||  Library of Mathematics, developed at the Computer Science     
+    ||T||  Department, University of Bologna, Italy.                     
+    ||I||                                                                
+    ||T||  HELM is free software; you can redistribute it and/or         
+    ||A||  modify it under the terms of the GNU General Public License   
+    \   /  version 2 or (at your option) any later version.      
+     \ /   This software is distributed as is, NO WARRANTY.     
+      V_______________________________________________________________ *)
+
+(* $Id: orderings.ml 9869 2009-06-11 22:52:38Z denes $ *)
+
+  let eqP () = 
+    let r = 
+      OCic2NCic.reference_of_oxuri 
+       (UriManager.uri_of_string 
+         "cic:/matita/logic/equality/eq.ind#xpointer(1/1)")
+    in
+    NCic.Const r
+  ;;
+
+  let eq_ind () = 
+    let r = 
+      OCic2NCic.reference_of_oxuri 
+       (UriManager.uri_of_string 
+         "cic:/matita/logic/equality/eq_ind.con")
+    in
+    NCic.Const r
+  ;;
+
+  let eq_ind_r () = 
+    let r = 
+      OCic2NCic.reference_of_oxuri 
+       (UriManager.uri_of_string 
+         "cic:/matita/logic/equality/eq_elim_r.con")
+    in
+    NCic.Const r
+  ;;
+
+  let eq_refl () = 
+    let r = 
+      OCic2NCic.reference_of_oxuri 
+       (UriManager.uri_of_string 
+         "cic:/matita/logic/equality/eq.ind#xpointer(1/1/1)")
+    in
+    NCic.Const r
+  ;;
+
+  let extract lift vl t =
+    let rec pos i = function 
+      | [] -> raise Not_found 
+      | j :: tl when j <> i -> 1+ pos i tl
+      | _ -> 1
+    in
+    let vl_len = List.length vl in
+    let rec extract = function
+      | Terms.Leaf x -> NCicSubstitution.lift (vl_len+lift) x
+      | Terms.Var j -> 
+           (try NCic.Rel (pos j vl) with Not_found -> NCic.Implicit `Term) 
+      | Terms.Node l -> NCic.Appl (List.map extract l)
+    in
+      extract t
+  ;;
+
+   let mk_predicate hole_type amount ft p1 vl =
+    let rec aux t p = 
+      match p with
+      | [] -> NCic.Rel 1
+      | n::tl ->
+          match t with
+          | Terms.Leaf _ 
+          | Terms.Var _ -> 
+(*                prerr_endline ("term: " ^ ppfot ft);             *)
+               prerr_endline ("path: " ^ String.concat "," 
+                 (List.map string_of_int p1));
+               assert false
+          | Terms.Node l -> 
+              let l = 
+                HExtlib.list_mapi
+                  (fun t i -> 
+                    if i = n then aux t tl 
+                    else extract amount (0::vl) t)
+                  l
+              in            
+              NCic.Appl l
+    in
+      NCic.Lambda("x", hole_type, aux ft (List.rev p1))
+    ;;
+
+  let mk_proof (bag : NCic.term Terms.bag) mp steps =
+    let module Subst = FoSubst in
+    let position i l = 
+      let rec aux = function
+       | [] -> assert false
+       | (j,_) :: tl when i = j -> 1
+       | _ :: tl -> 1 + aux tl
+      in
+        aux l
+    in
+    let vars_of i l = fst (List.assoc i l) in
+    let ty_of i l = snd (List.assoc i l) in
+    let close_with_lambdas vl t = 
+      List.fold_left 
+       (fun t i -> 
+          NCic.Lambda ("x"^string_of_int i, NCic.Implicit `Type, t))
+       t vl  
+    in
+    let close_with_forall vl t = 
+      List.fold_left 
+       (fun t i -> 
+          NCic.Prod ("x"^string_of_int i, NCic.Implicit `Type, t))
+       t vl  
+    in
+    let get_literal id =
+      let _, lit, vl, proof = Terms.M.find id bag in
+      let lit =match lit with 
+          | Terms.Predicate t -> assert false 
+          | Terms.Equation (l,r,ty,_) -> 
+              Terms.Node [ Terms.Leaf eqP(); ty; l; r]
+      in
+       lit, vl, proof
+    in
+    let rec aux ongoal seen = function
+      | [] -> NCic.Rel 1
+      | id :: tl ->
+          let amount = List.length seen in
+          let lit,vl,proof = get_literal id in
+          if not ongoal && id = mp then
+            ((*prerr_endline ("Reached m point, id=" ^ (string_of_int id));*)
+             NCic.LetIn ("clause_" ^ string_of_int id, 
+                extract amount [] lit, 
+                (NCic.Appl [eq_refl();NCic.Implicit `Type;NCic.Implicit `Term]),
+                aux true ((id,([],lit))::seen) (id::tl)))
+          else
+          match proof with
+          | Terms.Exact _ when tl=[] ->
+             (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id));*)
+             aux ongoal seen tl
+          | Terms.Step _ when tl=[] -> assert false
+          | Terms.Exact ft ->
+             (* prerr_endline ("Exact for " ^ (string_of_int id));*)
+               NCic.LetIn ("clause_" ^ string_of_int id, 
+                 close_with_forall vl (extract amount vl lit),
+                 close_with_lambdas vl (extract amount vl ft),
+                 aux ongoal 
+                   ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
+          | Terms.Step (_, id1, id2, dir, pos, subst) ->
+              let id, id1,(lit,vl,proof) =
+               if ongoal then id1,id,get_literal id1
+               else id,id1,(lit,vl,proof)
+             in
+             let vl = if ongoal then Subst.filter subst vl else vl in
+              let proof_of_id id = 
+                let vars = List.rev (vars_of id seen) in
+                let args = List.map (Subst.apply_subst subst) vars in
+                let args = List.map (extract amount vl) args in
+               let rel_for_id = NCic.Rel (List.length vl + position id seen) in
+                 if args = [] then rel_for_id              
+                  else NCic.Appl (rel_for_id::args)
+              in
+              let p_id1 = proof_of_id id1 in
+              let p_id2 = proof_of_id id2 in
+              let pred, hole_type, l, r = 
+                let id1_ty = ty_of id1 seen in
+                let id2_ty,l,r = 
+                  match ty_of id2 seen with
+                  | Terms.Node [ _; t; l; r ] -> 
+                      extract amount vl (Subst.apply_subst subst t),
+                      extract amount vl (Subst.apply_subst subst l),
+                      extract amount vl (Subst.apply_subst subst r)
+                  | _ -> assert false
+                in
+                 (*prerr_endline "mk_predicate :";
+                 if ongoal then prerr_endline "ongoal=true"
+                 else prerr_endline "ongoal=false";
+                 prerr_endline ("id=" ^ string_of_int id);
+                 prerr_endline ("id1=" ^ string_of_int id1);
+                 prerr_endline ("id2=" ^ string_of_int id2);
+                 prerr_endline ("Positions :" ^
+                                  (String.concat ", "
+                                     (List.map string_of_int pos)));*)
+                mk_predicate 
+                  id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
+                id2_ty,
+                l,r
+              in
+              let l, r, eq_ind = 
+                if (ongoal=true) = (dir=Terms.Left2Right) then
+                  r,l,eq_ind_r ()
+                else
+                  l,r,eq_ind ()
+              in
+               NCic.LetIn ("clause_" ^ string_of_int id, 
+                 close_with_forall vl (extract amount vl lit),
+                          (* NCic.Implicit `Type, *)
+                 close_with_lambdas vl 
+                   (NCic.Appl [ eq_ind ; hole_type; l; pred; p_id1; r; p_id2 ]),
+                 aux ongoal 
+                   ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
+    in 
+      aux false [] steps
+  ;;
+
diff --git a/helm/software/components/ng_paramodulation/nCicProof.mli b/helm/software/components/ng_paramodulation/nCicProof.mli
new file mode 100644 (file)
index 0000000..7a10324
--- /dev/null
@@ -0,0 +1,17 @@
+(*
+    ||M||  This file is part of HELM, an Hypertextual, Electronic        
+    ||A||  Library of Mathematics, developed at the Computer Science     
+    ||T||  Department, University of Bologna, Italy.                     
+    ||I||                                                                
+    ||T||  HELM is free software; you can redistribute it and/or         
+    ||A||  modify it under the terms of the GNU General Public License   
+    \   /  version 2 or (at your option) any later version.      
+     \ /   This software is distributed as is, NO WARRANTY.     
+      V_______________________________________________________________ *)
+
+(* $Id: nCic.ml 9058 2008-10-13 17:42:30Z tassi $ *)
+
+val mk_proof: 
+      NCic.term Terms.bag -> Terms.M.key -> Terms.M.key list -> NCic.term
+
+
index 1e76758e5b08ead174a79f660761d0954758724c..b898d353babce559769c77239d630d1bae9b404b 100644 (file)
@@ -1,80 +1,95 @@
-let debug s =
-   () (*prerr_endline s*)
-;;
+(*
+    ||M||  This file is part of HELM, an Hypertextual, Electronic        
+    ||A||  Library of Mathematics, developed at the Computer Science     
+    ||T||  Department, University of Bologna, Italy.                     
+    ||I||                                                                
+    ||T||  HELM is free software; you can redistribute it and/or         
+    ||A||  modify it under the terms of the GNU General Public License   
+    \   /  version 2 or (at your option) any later version.      
+     \ /   This software is distributed as is, NO WARRANTY.     
+      V_______________________________________________________________ *)
 
-let nparamod rdb metasenv subst context t table =
-  let max_nb_iter = 999999999 in
-  let amount_of_time = 300.0 in
-  let module C = struct
-    let metasenv = metasenv
-    let subst = subst
-    let context = context
-  end
-  in
-  let nb_iter = ref 0 in
-  let module B = NCicBlob.NCicBlob(C) in
-  let module Pp = Pp.Pp (B) in
-  let module FU = FoUnif.Founif(B) in
-  let module IDX = Index.Index(B) in
-  let module Sup = Superposition.Superposition(B) in
-  let module Utils = FoUtils.Utils(B) in
+(* $Id: orderings.ml 9869 2009-06-11 22:52:38Z denes $ *)
+
+let debug s = prerr_endline s ;;
+(* let debug _ = ();; *)
     
-  let module WeightOrderedPassives =
+let max_nb_iter = 999999999 ;;
+let amount_of_time = 300.0 ;;
+
+module Paramod (B : Terms.Blob) = struct
+  exception Failure of string * B.t Terms.bag * int * int
+  type bag = B.t Terms.bag * int
+  module Pp = Pp.Pp (B) 
+  module FU = FoUnif.Founif(B) 
+  module IDX = Index.Index(B) 
+  module Sup = Superposition.Superposition(B) 
+  module Utils = FoUtils.Utils(B) 
+  module WeightOrderedPassives =
       struct
        type t = B.t Terms.passive_clause
-           
        let compare = Utils.compare_passive_clauses_weight
       end
-  in
-  let module AgeOrderedPassives =
+
+  module AgeOrderedPassives =
       struct
        type t = B.t Terms.passive_clause
-           
        let compare = Utils.compare_passive_clauses_age
       end
-  in
-  let module WeightPassiveSet = Set.Make(WeightOrderedPassives) in
-  let module AgePassiveSet = Set.Make(AgeOrderedPassives) in
+  
+  module WeightPassiveSet = Set.Make(WeightOrderedPassives)
+  module AgePassiveSet = Set.Make(AgeOrderedPassives)
+
   let add_passive_clause ?(no_weight=false) (passives_w,passives_a) cl =
     let cl = if no_weight then (0,cl)
     else Utils.mk_passive_clause cl in
     WeightPassiveSet.add cl passives_w, AgePassiveSet.add cl passives_a
-  in
+  ;;
+
   let remove_passive_clause (passives_w,passives_a) cl =
     let passives_w = WeightPassiveSet.remove cl passives_w in
     let passives_a = AgePassiveSet.remove cl passives_a in
       passives_w,passives_a
-  in
+  ;;
+
   let add_passive_clauses (passives_w,passives_a) new_clauses =
     let new_clauses_w,new_clauses_a = List.fold_left add_passive_clause
       (WeightPassiveSet.empty,AgePassiveSet.empty) new_clauses
     in
       (WeightPassiveSet.union new_clauses_w passives_w,
        AgePassiveSet.union new_clauses_a passives_a)
-  in
+  ;;
+
   let is_passive_set_empty (passives_w,passives_a) =
     if (WeightPassiveSet.is_empty passives_w) then begin
       assert (AgePassiveSet.is_empty passives_a); true
     end else begin
       assert (not (AgePassiveSet.is_empty passives_a)); false
     end
-  in
-  let passive_set_cardinal (passives_w,_) =
-    WeightPassiveSet.cardinal passives_w
-  in
-  let passive_singleton cl =
-    (WeightPassiveSet.singleton cl,AgePassiveSet.singleton cl)
-  in
+  ;;
+
+  let passive_set_cardinal (passives_w,_) = WeightPassiveSet.cardinal passives_w
+  
   let passive_empty_set =
     (WeightPassiveSet.empty,AgePassiveSet.empty)
-  in
+  ;;
+
   let pick_min_passive use_age (passives_w,passives_a) =
     if use_age then AgePassiveSet.min_elt passives_a
     else WeightPassiveSet.min_elt passives_w
-  in
-  let timeout = Unix.gettimeofday () +. amount_of_time in 
+  ;;
 
-    (* TODO : global age over facts and goals (without comparing weights) *)
+  let mk_clause bag maxvar (t,ty) =
+    let (proof,ty) = B.saturate t ty in
+    let c, maxvar = Utils.mk_unit_clause maxvar ty proof in
+    let bag, c = Utils.add_to_bag bag c in
+    (bag, maxvar), c
+  ;;
+  
+  let mk_passive (bag,maxvar) = mk_clause bag maxvar;;
+  let mk_goal (bag,maxvar) = mk_clause bag maxvar;;
+
+  (* TODO : global age over facts and goals (without comparing weights) *)
   let select passives g_passives =
     if is_passive_set_empty passives then begin
       assert (not (is_passive_set_empty g_passives));
@@ -90,7 +105,7 @@ let nparamod rdb metasenv subst context t table =
            (false,snd cl,remove_passive_clause passives cl,g_passives)
          else
            (true,snd g_cl,passives,remove_passive_clause g_passives g_cl)
-  in
+  ;;
 
   let backward_infer_step bag maxvar actives passives
                           g_actives g_passives g_current =
@@ -102,7 +117,7 @@ let nparamod rdb metasenv subst context t table =
        debug "Performed infer_left step";
         bag, maxvar, actives, passives, g_current::g_actives,
     (add_passive_clauses g_passives new_goals)
-  in
+  ;;
 
   let forward_infer_step bag maxvar actives passives g_actives
                          g_passives current =
@@ -145,18 +160,18 @@ let nparamod rdb metasenv subst context t table =
        bag, maxvar, actives,
     add_passive_clauses passives new_clauses, g_actives,
     add_passive_clauses g_passives new_goals
-  in
-
-  let rec given_clause bag maxvar actives passives g_actives g_passives =
+  ;;
+  let rec given_clause bag maxvar nb_iter timeout actives passives g_actives g_passives =
     (* prerr_endline "Bag :"; prerr_endline (Pp.pp_bag bag);
       prerr_endline "Active table :"; 
        (List.iter (fun x -> prerr_endline (Pp.pp_unit_clause x))
          (fst actives)); *)
-    incr nb_iter; if !nb_iter = max_nb_iter then       
-      raise (Failure "No iterations left !");
+    let nb_iter = nb_iter + 1 in
+    if nb_iter = max_nb_iter then       
+      raise (Failure ("No iterations left !",bag,maxvar,nb_iter));
     if Unix.gettimeofday () > timeout then
-      raise (Failure "Timeout !");
-
+      raise (Failure ("Timeout !",bag,maxvar,nb_iter));
 
     let rec aux_select passives g_passives =
       let backward,current,passives,g_passives = select passives g_passives in
@@ -190,77 +205,49 @@ let nparamod rdb metasenv subst context t table =
       debug
        (Printf.sprintf "Number of passives : %d"
           (passive_set_cardinal passives));
-      given_clause bag maxvar actives passives g_actives g_passives
-  in
+      given_clause 
+        bag maxvar nb_iter timeout actives passives g_actives g_passives
+  ;;
 
-  let mk_clause bag maxvar (t,ty) =
-    let (proof,ty) = B.saturate t ty in
-    let c, maxvar = Utils.mk_unit_clause maxvar ty proof in
-    let bag, c = Utils.add_to_bag bag c in
-    bag, maxvar, c
-  in
-  let bag, maxvar, goal = mk_clause Terms.M.empty 0 t in
-  let g_actives = [] in
-  let g_passives =
-    (* passive_singleton (Utils.mk_passive_clause goal)*)
-    passive_singleton (0,goal)
-  in
-  let passives, bag, maxvar = 
-    List.fold_left
-      (fun (cl, bag, maxvar) t -> 
-         let bag, maxvar, c = mk_clause bag maxvar t in
-         (add_passive_clause ~no_weight:true cl c), bag, maxvar)
-      (passive_empty_set, bag, maxvar) table 
-  in
-  let actives = [], IDX.DT.empty in
-  try given_clause bag maxvar actives passives g_actives g_passives
-  with 
-  | Sup.Success (bag, _, (i,_,_,_)) ->
-      let l =
-       let rec traverse ongoal (accg,acce) i =
-         match Terms.M.find i bag with
-           | (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,_,_,_)) ->
-               if (not ongoal) && (List.mem i acce) then accg,acce
-               else
-                  let accg,acce = 
-                   traverse false (traverse ongoal (accg,acce) i1) i2
-                 in
-                   if ongoal then i::accg,acce else accg,i::acce
-        in
-       let gsteps,esteps = traverse true ([],[]) i in
-         (List.rev esteps)@gsteps
-      in
-       prerr_endline (Printf.sprintf "Found proof in %d iterations, %fs"
-                        !nb_iter
-                        (Unix.gettimeofday() -. timeout +. amount_of_time));
-      (* prerr_endline "Proof:"; 
-      List.iter (fun x -> prerr_endline (string_of_int x);
-              prerr_endline (Pp.pp_unit_clause (Terms.M.find x bag))) l;*)
-      let proofterm = B.mk_proof bag i l in
-       prerr_endline (Printf.sprintf "Got proof term, %fs"
-                        (Unix.gettimeofday() -. timeout +. amount_of_time));
-      let metasenv, proofterm = 
-        let rec aux k metasenv = function
-          | NCic.Meta _ as t -> metasenv, t
-          | NCic.Implicit _ -> 
-              let metasenv,i,_,_=NCicMetaSubst.mk_meta metasenv context `Term in
-              metasenv, NCic.Meta (i,(k,NCic.Irl (List.length context)))
-          | t -> NCicUntrusted.map_term_fold_a 
-                  (fun _ k -> k+1) k aux metasenv t
+  let paramod (bag,maxvar) ~g_passives ~passives =
+    let timeout = Unix.gettimeofday () +. amount_of_time in 
+    let passives = add_passive_clauses passive_empty_set passives in
+    let g_passives = add_passive_clauses passive_empty_set g_passives in
+    let g_actives = [] in
+    let actives = [], IDX.DT.empty in
+    try 
+     given_clause 
+      bag maxvar 0 timeout actives passives g_actives g_passives
+    with 
+    | Sup.Success (bag, _, (i,_,_,_)) ->
+        let l =
+          let rec traverse ongoal (accg,acce) i =
+            match Terms.M.find i bag with
+              | (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,_,_,_)) ->
+               if (not ongoal) && (List.mem i acce) then accg,acce
+               else
+                    let accg,acce = 
+                   traverse false (traverse ongoal (accg,acce) i1) i2
+                 in
+                   if ongoal then i::accg,acce else accg,i::acce
+          in
+          let gsteps,esteps = traverse true ([],[]) i in
+            (List.rev esteps)@gsteps
         in
-         aux 0 metasenv proofterm
-      in
-      let metasenv, subst, proofterm, _prooftype = 
-        NCicRefiner.typeof 
-          (rdb#set_coerc_db NCicCoercion.empty_db) 
-          metasenv subst context proofterm None
-      in
-      proofterm, metasenv, subst
-  | Failure s -> 
-      prerr_endline s;
-      prerr_endline
-      (Printf.sprintf "FAILURE in %d iterations" !nb_iter); assert false
-;;
+        prerr_endline 
+          (Printf.sprintf "Found proof, %fs" 
+            (Unix.gettimeofday() -. timeout +. amount_of_time));
+        prerr_endline "Proof:"; 
+        List.iter (fun x -> prerr_endline (string_of_int x);
+                prerr_endline (Pp.pp_unit_clause (Terms.M.find x bag))) l;
+        [ bag, i, l ]
+    | Failure (msg,_bag,_maxvar,nb_iter) -> 
+        prerr_endline msg;
+        prerr_endline (Printf.sprintf "FAILURE in %d iterations" nb_iter); 
+        []
+  ;;
+
+end
index 821ac0e2cbb8f08ef0255ec3e6c47771df2bebdb..2664494993095eac6ea96f1464638ffd51930a7a 100644 (file)
@@ -1,5 +1,23 @@
-val nparamod : 
-  #NRstatus.status ->
-  NCic.metasenv -> NCic.substitution -> NCic.context -> 
-    (NCic.term * NCic.term) -> (NCic.term * NCic.term) list ->
-     NCic.term * NCic.metasenv * NCic.substitution
+(*
+    ||M||  This file is part of HELM, an Hypertextual, Electronic        
+    ||A||  Library of Mathematics, developed at the Computer Science     
+    ||T||  Department, University of Bologna, Italy.                     
+    ||I||                                                                
+    ||T||  HELM is free software; you can redistribute it and/or         
+    ||A||  modify it under the terms of the GNU General Public License   
+    \   /  version 2 or (at your option) any later version.      
+     \ /   This software is distributed as is, NO WARRANTY.     
+      V_______________________________________________________________ *)
+
+(* $Id: orderings.ml 9869 2009-06-11 22:52:38Z denes $ *)
+
+module Paramod ( B : Terms.Blob ) : 
+  sig
+    type bag = B.t Terms.bag * int
+    val mk_passive : bag -> B.t * B.t -> bag * B.t Terms.unit_clause
+    val mk_goal : bag -> B.t * B.t -> bag * B.t Terms.unit_clause
+    val paramod : 
+      bag -> g_passives:B.t Terms.unit_clause list -> 
+      passives:B.t Terms.unit_clause list ->
+       (B.t Terms.bag * int * int list) list
+  end
index 0118df7c18ea4f3fc2cdcd17aaad1a2f2daeaa82..bcc3bcf7bd8a05ff4906f35abafcd6491de67eb2 100644 (file)
@@ -67,6 +67,5 @@ module type Blob =
     val pp : t -> string
     val embed : t -> t foterm
     val saturate : t -> t -> t foterm * t foterm
-    val mk_proof : t bag -> int -> int list -> t
   end
 
index f1eacb7346ae998bdb7e7ba8afe8da46e410cd69..c9c99caa1dfb8295474d32a185dd5f79181945a3 100644 (file)
@@ -75,7 +75,5 @@ module type Blob =
     (* saturate [proof] [type] -> [proof] * [type] *)
     val saturate : t -> t -> t foterm * t foterm
 
-    val mk_proof : t bag -> int -> int list -> t
-
   end
 
index fc0a8013cca3b09deb1b8ab9703499dc0f66d4f5..a384d54d56259f89952a7a14f8d013022cee568a 100644 (file)
@@ -601,11 +601,13 @@ let auto ~params:(l,_) status goal =
         (status, (t,ty) :: l))
       (status,[]) l
   in
-  let pt, metasenv, subst = 
-    Paramod.nparamod status metasenv subst (ctx_of gty) (NCic.Rel ~-1,t) l 
-  in      
-  let status = status#set_obj (n,h,metasenv,subst,o) in
-  instantiate status goal (NTacStatus.mk_cic_term (ctx_of gty) pt)
+  match
+    NCicParamod.nparamod status metasenv subst (ctx_of gty) (NCic.Rel ~-1,t) l 
+  with
+  | [] -> raise (NTacStatus.Error (lazy "no proof found",None))
+  | (pt, metasenv, subst)::_ -> 
+      let status = status#set_obj (n,h,metasenv,subst,o) in
+      instantiate status goal (NTacStatus.mk_cic_term (ctx_of gty) pt)
 ;;
 
 let auto_tac ~params status =