]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/components/ng_paramodulation/nCicProof.ml
Use of standard OCaml syntax
[helm.git] / matita / components / ng_paramodulation / nCicProof.ml
index 3f30a85c30ca84422d9b7554f3050581a0d4ac92..706131f8bc73eb64db0d89f02b31a4d6b760234f 100644 (file)
@@ -64,7 +64,6 @@ let debug c _ = c;;
       extract t
   ;;
 
-
    let mk_predicate status hole_type amount ft p1 vl =
     let rec aux t p = 
       match p with
@@ -73,16 +72,16 @@ let debug c _ = c;;
           match t with
           | Terms.Leaf _ 
           | Terms.Var _ -> 
-              let module NCicBlob = NCicBlob.NCicBlob(
+              (*let module NCicBlob = NCicBlob.NCicBlob(
                         struct
                           let metasenv = [] let subst = [] let context = []
                         end)
                           in
-              let module Pp = Pp.Pp(NCicBlob) in  
-               prerr_endline ("term: " ^ Pp.pp_foterm ft);
-               prerr_endline ("path: " ^ String.concat "," 
+               let module Pp = Pp.Pp(NCicBlob) in  
+                  prerr_endline ("term: " ^ Pp.pp_foterm ft); 
+                  prerr_endline ("path: " ^ String.concat "," 
                  (List.map string_of_int p1));
-               prerr_endline ("leading to: " ^ Pp.pp_foterm t);
+               prerr_endline ("leading to: " ^ Pp.pp_foterm t); *)
                assert false
           | Terms.Node l -> 
               let l = 
@@ -132,7 +131,7 @@ let debug c _ = c;;
      NCic.Implicit `Term; eq; NCic.Implicit `Term]; 
   ;;
 
-  let trans eq p = 
+  let trans eq _p = 
    let u= NUri.uri_of_string "cic:/matita/ng/properties/relations/trans.con" in
    let u = NReference.reference_of_spec u (NReference.Fix(0,1,3)) in
      NCic.Appl[NCic.Const u; NCic.Implicit `Type; NCic.Implicit `Term;
@@ -181,7 +180,7 @@ let debug c _ = c;;
               List.fold_left
                (fun (i,acc) t ->
                  i+1,
-                  let f = extract status amount vl f in
+                  let _f = extract status amount vl f in
                   if i = n then
                    let imp = NCic.Implicit `Term in
                     NCic.Appl (dag::imp::imp::imp(* f *)::imp::imp::
@@ -193,19 +192,23 @@ let debug c _ = c;;
     ;;
 
   let mk_proof status ?(demod=false) (bag : NCic.term Terms.bag) mp subst steps=
-    let module NCicBlob = 
+    (*let module NCicBlob = 
        NCicBlob.NCicBlob(
         struct
           let metasenv = [] let subst = [] let context = []
         end)
      in
-     let  module Pp = Pp.Pp(NCicBlob) 
-     in
+     let  module Pp = Pp.Pp(NCicBlob) in*)
     let module Subst = FoSubst in
+    (*let compose subst1 subst2 =
+      let s1 = List.map (fun (x,t) -> (x, Subst.apply_subst subst2 t)) subst1 in
+      let s2 = List.filter (fun (m,_) -> not (List.mem_assoc m s1)) subst2
+      in s1 @ s2
+    in*)
     let position i l = 
       let rec aux = function
        | [] -> assert false
-       | (j,_) :: tl when i = j -> 1
+       | (j,_) :: _ when i = j -> 1
        | _ :: tl -> 1 + aux tl
       in
         aux l
@@ -229,7 +232,7 @@ let debug c _ = c;;
       let lit =match lit with 
           | Terms.Predicate t -> t (* assert false *) 
           | Terms.Equation (l,r,ty,_) -> 
-              Terms.Node [ Terms.Leaf eqP(); ty; l; r]
+              Terms.Node [ Terms.Leaf (eqP()); ty; l; r]
       in
         lit, vl, proof
     in
@@ -238,24 +241,7 @@ let debug c _ = c;;
       let lit = Subst.apply_subst subst lit in
         extract status 0 [] lit in
     (* composition of all subst acting on goal *)
-    let res_subst =
-      let rec rsaux ongoal acc = 
-       function
-         | [] -> acc (* is the final subst for refl *)
-         | id::tl when ongoal ->
-            let lit,vl,proof = get_literal id in
-             (match proof with
-               | Terms.Exact _ -> rsaux ongoal acc tl
-               | Terms.Step (_, _, _, _, _, s) ->
-                   rsaux ongoal (s@acc) tl)
-         | id::tl -> 
-           (* subst is the the substitution for refl *)
-           rsaux (id=mp) subst tl
-      in 
-      let r = rsaux false [] steps in
-       (* prerr_endline ("res substitution: " ^ Pp.pp_substitution r);
-           prerr_endline ("\n" ^ "subst: " ^ Pp.pp_substitution subst); *)
-      r
+    let res_subst = subst
     in
     let rec aux ongoal seen = function
       | [] -> NCic.Rel 1
@@ -268,16 +254,17 @@ let debug c _ = c;;
             let refl = 
              if demod then NCic.Implicit `Term 
              else mk_refl eq_ty in
-             (* prerr_endline ("Reached m point, id=" ^ (string_of_int id));
+              (* prerr_endline ("Reached m point, id=" ^ (string_of_int id));
                 (NCic.LetIn ("clause_" ^ string_of_int id, eq_ty, refl,
                 aux true ((id,([],lit))::seen) (id::tl))) *)
-              NCicSubstitution.subst status
+              let res = NCicSubstitution.subst status
                 ~avoid_beta_redexes:true ~no_implicit:false refl
                 (aux true ((id,([],lit))::seen) (id::tl))
+              in res
           else
           match proof with
           | Terms.Exact _ when tl=[] ->
-              (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id));*)
+              (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id)); *)
               aux ongoal seen tl
           | Terms.Step _ when tl=[] -> assert false
           | Terms.Exact ft ->
@@ -288,21 +275,27 @@ let debug c _ = c;;
                  close_with_lambdas vl (extract status amount vl ft),
                  aux ongoal 
                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
-              *)
-               NCicSubstitution.subst status
+               *)
+               let res = NCicSubstitution.subst status
                  ~avoid_beta_redexes:true ~no_implicit:false
                  (close_with_lambdas vl (extract status amount vl ft))
                  (aux ongoal 
                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
+               in res 
           | Terms.Step (_, id1, id2, dir, pos, subst) ->
-              let id, id1,(lit,vl,proof) =
+              (* prerr_endline (if ongoal then "on goal" else "not on goal");
+                 prerr_endline (Pp.pp_substitution subst); *)
+              (* let subst = if ongoal then res_subst else subst in *)
+              let id, id1,(lit,vl,_proof) =
                 if ongoal then
                  let lit,vl,proof = get_literal id1 in
-                 id1,id,(Subst.apply_subst res_subst lit, 
-                         Subst.filter res_subst vl, proof)
+                 id1,id,(Subst.apply_subst res_subst lit,
+                          Subst.filter res_subst vl, proof)
                 else id,id1,(lit,vl,proof) in
               (* free variables remaining in the goal should not
                  be abstracted: we do not want to prove a generalization *)
+              (* prerr_endline ("variables = " ^ String.concat ","
+                (List.map string_of_int vl)); *)
               let vl = if ongoal then [] else vl in 
               let proof_of_id id = 
                 let vars = List.rev (vars_of id seen) in
@@ -314,37 +307,6 @@ let debug c _ = c;;
               in
               let p_id1 = proof_of_id id1 in
               let p_id2 = proof_of_id id2 in
-(*
-              let morphism, l, r =
-                let p =                
-                 if (ongoal=true) = (dir=Terms.Left2Right) then
-                   p_id2 
-                 else sym p_id2 in
-                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 status amount vl (Subst.apply_subst subst t),
-                      extract status amount vl (Subst.apply_subst subst l),
-                      extract status 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_morphism
-                  p amount (Subst.apply_subst subst id1_ty) pos vl,
-                l,r
-              in
-              let rewrite_step = iff1 morphism p_id1
-             in
-*)
               let pred, hole_type, l, r = 
                 let id1_ty = ty_of id1 seen in
                 let id2_ty,l,r = 
@@ -355,19 +317,9 @@ let debug c _ = c;;
                       extract status 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 
-                                ^": " ^ Pp.pp_foterm id1_ty);
-                  prerr_endline ("id2=" ^ string_of_int id2
-                                ^ ": " ^ NCicPp.ppterm [][][] id2_ty);
-                  prerr_endline ("Positions :" ^
-                                   (String.concat ", "
-                                      (List.map string_of_int pos)));*)
-                mk_predicate status
+                (* let _ = prerr_endline ("body = " ^(Pp.pp_foterm id1_ty)) 
+                   in *)
+               mk_predicate status
                   id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
                 id2_ty,
                 l,r