]> matita.cs.unibo.it Git - helm.git/blob - helm/software/components/ng_paramodulation/nCicProof.ml
59d7e95ff65f824b0e70b968aa4b16d1ecbb9aec
[helm.git] / helm / software / components / ng_paramodulation / nCicProof.ml
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic        
3     ||A||  Library of Mathematics, developed at the Computer Science     
4     ||T||  Department, University of Bologna, Italy.                     
5     ||I||                                                                
6     ||T||  HELM is free software; you can redistribute it and/or         
7     ||A||  modify it under the terms of the GNU General Public License   
8     \   /  version 2 or (at your option) any later version.      
9      \ /   This software is distributed as is, NO WARRANTY.     
10       V_______________________________________________________________ *)
11
12 (* $Id: orderings.ml 9869 2009-06-11 22:52:38Z denes $ *)
13
14 type eq_sig_type = Eq | EqInd_l | EqInd_r | Refl
15
16 let eqsig = ref (fun _ -> assert false);;
17 let set_sig f = eqsig:= 
18 f;;
19
20
21 let default_sig = function
22   | Eq -> 
23       let uri = NUri.uri_of_string "cic:/matita/ng/Plogic/equality/peq.ind" in
24       let ref = NReference.reference_of_spec uri (NReference.Ind(true,0,2)) in
25         NCic.Const ref
26   | EqInd_l -> 
27       let uri = NUri.uri_of_string "cic:/matita/ng/Plogic/equality/rewrite_l.con" in
28       let ref = NReference.reference_of_spec uri (NReference.Def(1)) in
29         NCic.Const ref
30   | EqInd_r -> 
31       let uri = NUri.uri_of_string "cic:/matita/ng/Plogic/equality/rewrite_r.con" in
32       let ref = NReference.reference_of_spec uri (NReference.Def(3)) in
33         NCic.Const ref
34   | Refl ->
35       let uri = NUri.uri_of_string "cic:/matita/ng/Plogic/equality/peq.ind" in
36       let ref = NReference.reference_of_spec uri (NReference.Con(0,1,2)) in
37         NCic.Const ref
38
39 let set_default_sig () =
40   prerr_endline "setting default sig";
41   eqsig := default_sig
42
43 let set_reference_of_oxuri reference_of_oxuri = 
44   prerr_endline "setting oxuri in nCicProof";
45   let nsig = function
46     | Eq -> 
47         NCic.Const
48           (reference_of_oxuri 
49              (UriManager.uri_of_string 
50                 "cic:/matita/logic/equality/eq.ind#xpointer(1/1)"))  
51     | EqInd_l -> 
52         NCic.Const
53           (reference_of_oxuri 
54              (UriManager.uri_of_string 
55                 "cic:/matita/logic/equality/eq_ind.con"))
56     | EqInd_r -> 
57         NCic.Const
58           (reference_of_oxuri 
59              (UriManager.uri_of_string 
60                 "cic:/matita/logic/equality/eq_elim_r.con"))
61     | Refl ->
62         NCic.Const
63           (reference_of_oxuri 
64              (UriManager.uri_of_string 
65                 "cic:/matita/logic/equality/eq.ind#xpointer(1/1/1)"))
66   in eqsig:= nsig
67   ;;
68
69 let debug c r = prerr_endline r; c 
70
71   let eqP() = prerr_endline "1"; prerr_endline "1"; debug (!eqsig Eq) "eqp"  ;;
72   let eq_ind() = prerr_endline "2"; debug (!eqsig EqInd_l) "eq_ind" ;;
73   let eq_ind_r() = prerr_endline "3"; debug (!eqsig EqInd_r) "eq_ind_r";; 
74   let eq_refl() = prerr_endline "4"; debug (!eqsig Refl) "refl";;
75
76
77   let extract lift vl t =
78     let rec pos i = function 
79       | [] -> raise Not_found 
80       | j :: tl when j <> i -> 1+ pos i tl
81       | _ -> 1
82     in
83     let vl_len = List.length vl in
84     let rec extract = function
85       | Terms.Leaf x -> NCicSubstitution.lift (vl_len+lift) x
86       | Terms.Var j -> 
87            (try NCic.Rel (pos j vl) with Not_found -> NCic.Implicit `Term) 
88       | Terms.Node l -> NCic.Appl (List.map extract l)
89     in
90       extract t
91   ;;
92
93    let mk_predicate hole_type amount ft p1 vl =
94     let rec aux t p = 
95       match p with
96       | [] -> NCic.Rel 1
97       | n::tl ->
98           match t with
99           | Terms.Leaf _ 
100           | Terms.Var _ -> 
101               let module Pp = 
102                 Pp.Pp(NCicBlob.NCicBlob(
103                         struct
104                           let metasenv = [] let subst = [] let context = []
105                         end))
106               in  
107                prerr_endline ("term: " ^ Pp.pp_foterm ft);
108                prerr_endline ("path: " ^ String.concat "," 
109                  (List.map string_of_int p1));
110                prerr_endline ("leading to: " ^ Pp.pp_foterm t);
111                assert false
112           | Terms.Node l -> 
113               let l = 
114                 HExtlib.list_mapi
115                   (fun t i -> 
116                     if i = n then aux t tl 
117                     else extract amount (0::vl) t)
118                   l
119               in            
120               NCic.Appl l
121     in
122       NCic.Lambda("x", hole_type, aux ft (List.rev p1))
123     ;;
124
125   let mk_proof (bag : NCic.term Terms.bag) mp steps =
126     let module Subst = FoSubst in
127     let position i l = 
128       let rec aux = function
129        | [] -> assert false
130        | (j,_) :: tl when i = j -> 1
131        | _ :: tl -> 1 + aux tl
132       in
133         aux l
134     in
135     let vars_of i l = fst (List.assoc i l) in
136     let ty_of i l = snd (List.assoc i l) in
137     let close_with_lambdas vl t = 
138       List.fold_left 
139        (fun t i -> 
140           NCic.Lambda ("x"^string_of_int i, NCic.Implicit `Type, t))
141        t vl  
142     in
143     let close_with_forall vl t = 
144       List.fold_left 
145        (fun t i -> 
146           NCic.Prod ("x"^string_of_int i, NCic.Implicit `Type, t))
147        t vl  
148     in
149     let get_literal id =
150       let (_, lit, vl, proof),_,_ = Terms.get_from_bag id bag in
151       let lit =match lit with 
152           | Terms.Predicate t -> assert false 
153           | Terms.Equation (l,r,ty,_) -> 
154               Terms.Node [ Terms.Leaf eqP(); ty; l; r]
155       in
156         lit, vl, proof
157     in
158     let rec aux ongoal seen = function
159       | [] -> NCic.Rel 1
160       | id :: tl ->
161           let amount = List.length seen in
162           let lit,vl,proof = get_literal id in
163           if not ongoal && id = mp then
164             ((*prerr_endline ("Reached m point, id=" ^ (string_of_int id));*)
165              NCic.LetIn ("clause_" ^ string_of_int id, 
166                 extract amount [] lit, 
167                 (NCic.Appl [eq_refl();NCic.Implicit `Type;NCic.Implicit `Term]),
168                 aux true ((id,([],lit))::seen) (id::tl)))
169           else
170           match proof with
171           | Terms.Exact _ when tl=[] ->
172               (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id));*)
173               aux ongoal seen tl
174           | Terms.Step _ when tl=[] -> assert false
175           | Terms.Exact ft ->
176               (* prerr_endline ("Exact for " ^ (string_of_int id));*)
177                NCic.LetIn ("clause_" ^ string_of_int id, 
178                  close_with_forall vl (extract amount vl lit),
179                  close_with_lambdas vl (extract amount vl ft),
180                  aux ongoal 
181                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
182           | Terms.Step (_, id1, id2, dir, pos, subst) ->
183               let id, id1,(lit,vl,proof) =
184                 if ongoal then id1,id,get_literal id1
185                 else id,id1,(lit,vl,proof)
186               in
187               let vl = if ongoal then [](*Subst.filter subst vl*) else vl in
188               let proof_of_id id = 
189                 let vars = List.rev (vars_of id seen) in
190                 let args = List.map (Subst.apply_subst subst) vars in
191                 let args = List.map (extract amount vl) args in
192                 let rel_for_id = NCic.Rel (List.length vl + position id seen) in
193                   if args = [] then rel_for_id              
194                   else NCic.Appl (rel_for_id::args)
195               in
196               let p_id1 = proof_of_id id1 in
197               let p_id2 = proof_of_id id2 in
198               let pred, hole_type, l, r = 
199                 let id1_ty = ty_of id1 seen in
200                 let id2_ty,l,r = 
201                   match ty_of id2 seen with
202                   | Terms.Node [ _; t; l; r ] -> 
203                       extract amount vl (Subst.apply_subst subst t),
204                       extract amount vl (Subst.apply_subst subst l),
205                       extract amount vl (Subst.apply_subst subst r)
206                   | _ -> assert false
207                 in
208                   (*prerr_endline "mk_predicate :";
209                   if ongoal then prerr_endline "ongoal=true"
210                   else prerr_endline "ongoal=false";
211                   prerr_endline ("id=" ^ string_of_int id);
212                   prerr_endline ("id1=" ^ string_of_int id1);
213                   prerr_endline ("id2=" ^ string_of_int id2);
214                   prerr_endline ("Positions :" ^
215                                    (String.concat ", "
216                                       (List.map string_of_int pos)));*)
217                 mk_predicate 
218                   id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
219                 id2_ty,
220                 l,r
221               in
222               let l, r, eq_ind = 
223                 if (ongoal=true) = (dir=Terms.Left2Right) then
224                   r,l,eq_ind_r ()
225                 else
226                   l,r,eq_ind ()
227               in
228                NCic.LetIn ("clause_" ^ string_of_int id, 
229                  close_with_forall vl (extract amount vl lit),
230                            (* NCic.Implicit `Type, *)
231                  close_with_lambdas vl 
232                    (NCic.Appl [ eq_ind ; hole_type; l; pred; p_id1; r; p_id2 ]),
233                  aux ongoal 
234                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
235     in 
236       aux false [] steps
237   ;;
238