]> matita.cs.unibo.it Git - helm.git/blob - matita/components/ng_paramodulation/nCicProof.ml
60f2cbafc916ad2e808051ce4b3dd36dc4b0d1db
[helm.git] / matita / 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:= f;;
18 let get_sig = fun x -> !eqsig x;;
19
20 let default_sig = function
21   | Eq -> 
22       let uri = NUri.uri_of_string "cic:/matita/basics/logic/eq.ind" in
23       let ref = NReference.reference_of_spec uri (NReference.Ind(true,0,2)) in
24         NCic.Const ref
25   | EqInd_l -> 
26       let uri = NUri.uri_of_string "cic:/matita/basics/logic/rewrite_l.con" in
27       let ref = NReference.reference_of_spec uri (NReference.Def(1)) in
28         NCic.Const ref
29   | EqInd_r -> 
30       let uri = NUri.uri_of_string "cic:/matita/basics/logic/rewrite_r.con" in
31       let ref = NReference.reference_of_spec uri (NReference.Def(3)) in
32         NCic.Const ref
33   | Refl ->
34       let uri = NUri.uri_of_string "cic:/matita/basics/logic/eq.ind" in
35       let ref = NReference.reference_of_spec uri (NReference.Con(0,1,2)) in
36         NCic.Const ref
37
38 let set_default_sig () =
39   (*prerr_endline "setting default sig";*)
40   eqsig := default_sig
41
42 (* let debug c r = prerr_endline r; c *)
43 let debug c _ = c;;
44
45   let eqP() = debug (!eqsig Eq) "eq"  ;;
46   let eq_ind() = debug (!eqsig EqInd_l) "eq_ind" ;;
47   let eq_ind_r() = debug (!eqsig EqInd_r) "eq_ind_r";; 
48   let eq_refl() = debug (!eqsig Refl) "refl";;
49
50
51   let extract status lift vl t =
52     let rec pos i = function 
53       | [] -> raise Not_found 
54       | j :: tl when j <> i -> 1+ pos i tl
55       | _ -> 1
56     in
57     let vl_len = List.length vl in
58     let rec extract = function
59       | Terms.Leaf x -> NCicSubstitution.lift status (vl_len+lift) x
60       | Terms.Var j -> 
61            (try NCic.Rel (pos j vl) with Not_found -> NCic.Implicit `Term) 
62       | Terms.Node l -> NCic.Appl (List.map extract l)
63     in
64       extract t
65   ;;
66
67    let mk_predicate status hole_type amount ft p1 vl =
68     let rec aux t p = 
69       match p with
70       | [] -> NCic.Rel 1
71       | n::tl ->
72           match t with
73           | Terms.Leaf _ 
74           | Terms.Var _ -> 
75               let module NCicBlob = NCicBlob.NCicBlob(
76                         struct
77                           let metasenv = [] let subst = [] let context = []
78                         end)
79                           in
80               let module Pp = Pp.Pp(NCicBlob) in  
81                (* prerr_endline ("term: " ^ Pp.pp_foterm ft); 
82                   prerr_endline ("path: " ^ String.concat "," 
83                  (List.map string_of_int p1));
84                prerr_endline ("leading to: " ^ Pp.pp_foterm t); *)
85                assert false
86           | Terms.Node l -> 
87               let l = 
88                 HExtlib.list_mapi
89                   (fun t i -> 
90                     if i = n then aux t tl 
91                     else extract status amount (0::vl) t)
92                   l
93               in            
94               NCic.Appl l
95     in
96       NCic.Lambda("x", hole_type, aux ft (List.rev p1))
97     ;;
98
99   let dag = 
100    let uri = NUri.uri_of_string "cic:/matita/ng/sets/setoids/prop1.con" in
101    let ref = NReference.reference_of_spec uri (NReference.Fix(0,2,4)) in
102      NCic.Const ref
103   ;;
104
105   (*
106   let eq_setoid = 
107    let uri = NUri.uri_of_string "cic:/matita/ng/sets/setoids/eq.con" in
108    let ref = NReference.reference_of_spec uri (NReference.Fix(0,0,2)) in
109      NCic.Const ref
110   ;;
111   *)
112
113   let sym eq = 
114    let u= NUri.uri_of_string "cic:/matita/ng/properties/relations/sym.con" in
115    let u = NReference.reference_of_spec u (NReference.Fix(0,1,3)) in
116      NCic.Appl[NCic.Const u; NCic.Implicit `Type; NCic.Implicit `Term;
117      NCic.Implicit `Term; NCic.Implicit `Term; eq]; 
118   ;;
119
120   let eq_morphism1 eq = 
121    let u= NUri.uri_of_string "cic:/matita/ng/sets/setoids/eq_is_morphism1.con" in
122    let u = NReference.reference_of_spec u (NReference.Def 4) in
123      NCic.Appl[NCic.Const u; NCic.Implicit `Term; NCic.Implicit `Term;
124      NCic.Implicit `Term; NCic.Implicit `Term; eq]; 
125   ;;
126
127   let eq_morphism2 eq = 
128    let u= NUri.uri_of_string "cic:/matita/ng/sets/setoids/eq_is_morphism2.con" in
129    let u = NReference.reference_of_spec u (NReference.Def 4) in
130      NCic.Appl[NCic.Const u; NCic.Implicit `Term; NCic.Implicit `Term;
131      NCic.Implicit `Term; eq; NCic.Implicit `Term]; 
132   ;;
133
134   let trans eq p = 
135    let u= NUri.uri_of_string "cic:/matita/ng/properties/relations/trans.con" in
136    let u = NReference.reference_of_spec u (NReference.Fix(0,1,3)) in
137      NCic.Appl[NCic.Const u; NCic.Implicit `Type; NCic.Implicit `Term;
138      NCic.Implicit `Term; NCic.Implicit `Term; NCic.Implicit `Term; eq]
139   ;;
140
141   let iff1 eq p = 
142    let uri = NUri.uri_of_string "cic:/matita/ng/logic/connectives/if.con" in
143    let ref = NReference.reference_of_spec uri (NReference.Fix(0,2,1)) in
144      NCic.Appl[NCic.Const ref; NCic.Implicit `Type; NCic.Implicit `Type; 
145                eq; p]; 
146   ;;
147
148 (*
149   let mk_refl = function
150       | NCic.Appl [_; _; x; _] -> 
151    let uri= NUri.uri_of_string "cic:/matita/ng/properties/relations/refl.con" in
152    let ref = NReference.reference_of_spec uri (NReference.Fix(0,1,3)) in
153     NCic.Appl[NCic.Const ref; NCic.Implicit `Type; NCic.Implicit `Term;
154     NCic.Implicit `Term(*x*)]
155       | _ -> assert false
156 *)   
157
158   let mk_refl = function
159     | NCic.Appl [_; ty; l; _]
160       -> NCic.Appl [eq_refl();ty;l]
161     | _ -> assert false
162
163
164    let mk_morphism status eq amount ft pl vl =
165     let rec aux t p = 
166       match p with
167       | [] -> eq
168       | n::tl ->
169           prerr_endline (string_of_int n);
170           match t with
171           | Terms.Leaf _ 
172           | Terms.Var _ -> assert false
173           | Terms.Node [] -> assert false
174           | Terms.Node [ Terms.Leaf eqt ; _; l; r ]
175              when (eqP ()) = eqt ->
176                if n=2 then eq_morphism1 (aux l tl)
177                else eq_morphism2 (aux r tl)
178           | Terms.Node (f::l) ->
179              snd (
180               List.fold_left
181                (fun (i,acc) t ->
182                  i+1,
183                   let f = extract status amount vl f in
184                   if i = n then
185                    let imp = NCic.Implicit `Term in
186                     NCic.Appl (dag::imp::imp::imp(* f *)::imp::imp::
187                                [aux t tl])
188                   else
189                     NCicUntrusted.mk_appl acc [extract status amount vl t]
190                ) (1,extract status amount vl f) l)
191     in aux ft (List.rev pl)
192     ;;
193
194   let mk_proof status ?(demod=false) (bag : NCic.term Terms.bag) mp subst steps=
195     let module NCicBlob = 
196        NCicBlob.NCicBlob(
197          struct
198            let metasenv = [] let subst = [] let context = []
199          end)
200      in
201      let  module Pp = Pp.Pp(NCicBlob) 
202      in
203     let module Subst = FoSubst in
204     let compose subst1 subst2 =
205       let s1 = List.map (fun (x,t) -> (x, Subst.apply_subst subst2 t)) subst1 in
206       let s2 = List.filter (fun (m,_) -> not (List.mem_assoc m s1)) subst2
207       in s1 @ s2
208     in
209     let position i l = 
210       let rec aux = function
211        | [] -> assert false
212        | (j,_) :: tl when i = j -> 1
213        | _ :: tl -> 1 + aux tl
214       in
215         aux l
216     in
217     let vars_of i l = fst (List.assoc i l) in
218     let ty_of i l = snd (List.assoc i l) in
219     let close_with_lambdas vl t = 
220       List.fold_left 
221        (fun t i -> 
222           NCic.Lambda ("x"^string_of_int i, NCic.Implicit `Type, t))
223        t vl  
224     in
225     let close_with_forall vl t = 
226       List.fold_left 
227        (fun t i -> 
228           NCic.Prod ("x"^string_of_int i, NCic.Implicit `Type, t))
229        t vl  
230     in
231     let get_literal id =
232       let (_, lit, vl, proof),_,_ = Terms.get_from_bag id bag in
233       let lit =match lit with 
234           | Terms.Predicate t -> t (* assert false *) 
235           | Terms.Equation (l,r,ty,_) -> 
236               Terms.Node [ Terms.Leaf eqP(); ty; l; r]
237       in
238         lit, vl, proof
239     in
240     let proof_type =
241       let lit,_,_ = get_literal mp in
242       let lit = Subst.apply_subst subst lit in
243         extract status 0 [] lit in
244     (* composition of all subst acting on goal *)
245     let res_subst = subst
246     in
247     let rec aux ongoal seen = function
248       | [] -> NCic.Rel 1
249       | id :: tl ->
250           let amount = List.length seen in
251           let lit,vl,proof = get_literal id in
252           if not ongoal && id = mp then
253             let lit = Subst.apply_subst subst lit in 
254             let eq_ty = extract status amount [] lit in
255             let refl = 
256               if demod then NCic.Implicit `Term 
257               else mk_refl eq_ty in
258               (* prerr_endline ("Reached m point, id=" ^ (string_of_int id));
259                 (NCic.LetIn ("clause_" ^ string_of_int id, eq_ty, refl,
260                 aux true ((id,([],lit))::seen) (id::tl))) *)
261               let res = NCicSubstitution.subst status
262                 ~avoid_beta_redexes:true ~no_implicit:false refl
263                 (aux true ((id,([],lit))::seen) (id::tl))
264               in res
265           else
266           match proof with
267           | Terms.Exact _ when tl=[] ->
268               (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id)); *)
269               aux ongoal seen tl
270           | Terms.Step _ when tl=[] -> assert false
271           | Terms.Exact ft ->
272               (*
273                prerr_endline ("Exact for " ^ (string_of_int id));
274                NCic.LetIn ("clause_" ^ string_of_int id, 
275                  close_with_forall vl (extract status amount vl lit),
276                  close_with_lambdas vl (extract status amount vl ft),
277                  aux ongoal 
278                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
279                *)
280                let res = NCicSubstitution.subst status
281                  ~avoid_beta_redexes:true ~no_implicit:false
282                  (close_with_lambdas vl (extract status amount vl ft))
283                  (aux ongoal 
284                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
285                in res 
286           | Terms.Step (_, id1, id2, dir, pos, subst) ->
287               (* prerr_endline (if ongoal then "on goal" else "not on goal");
288                  prerr_endline (Pp.pp_substitution subst); *)
289               (* let subst = if ongoal then res_subst else subst in *)
290               let id, id1,(lit,vl,proof) =
291                 if ongoal then
292                   let lit,vl,proof = get_literal id1 in
293                   id1,id,(Subst.apply_subst res_subst lit,
294                           Subst.filter res_subst vl, proof)
295                 else id,id1,(lit,vl,proof) in
296               (* free variables remaining in the goal should not
297                  be abstracted: we do not want to prove a generalization *)
298               (* prerr_endline ("variables = " ^ String.concat ","
299                 (List.map string_of_int vl)); *)
300               let vl = if ongoal then [] else vl in 
301               let proof_of_id id = 
302                 let vars = List.rev (vars_of id seen) in
303                 let args = List.map (Subst.apply_subst subst) vars in
304                 let args = List.map (extract status amount vl) args in
305                 let rel_for_id = NCic.Rel (List.length vl + position id seen) in
306                   if args = [] then rel_for_id                    
307                   else NCic.Appl (rel_for_id::args)
308               in
309               let p_id1 = proof_of_id id1 in
310               let p_id2 = proof_of_id id2 in
311               let pred, hole_type, l, r = 
312                 let id1_ty = ty_of id1 seen in
313                 let id2_ty,l,r = 
314                   match ty_of id2 seen with
315                   | Terms.Node [ _; t; l; r ] -> 
316                       extract status amount vl (Subst.apply_subst subst t),
317                       extract status amount vl (Subst.apply_subst subst l),
318                       extract status amount vl (Subst.apply_subst subst r)
319                   | _ -> assert false
320                 in
321                 (* let _ = prerr_endline ("body = " ^(Pp.pp_foterm id1_ty)) 
322                    in *)
323                 mk_predicate status
324                   id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
325                 id2_ty,
326                 l,r
327               in
328               let rewrite_step =
329                 if (ongoal=true) = (dir=Terms.Left2Right) then
330                   NCic.Appl 
331                     [eq_ind_r(); hole_type; r; pred; p_id1; l; p_id2]
332                 else
333                   NCic.Appl 
334                     [ eq_ind(); hole_type; l; pred; p_id1; r; p_id2]
335               in
336               let body = aux ongoal 
337                 ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl 
338               in 
339               let occ =
340                NCicUntrusted.count_occurrences status [] 1 body in
341                 if occ <= 1 then
342                   NCicSubstitution.subst status
343                     ~avoid_beta_redexes:true ~no_implicit:false
344                     (close_with_lambdas vl rewrite_step) body
345                 else
346                   NCic.LetIn ("clause_" ^ string_of_int id, 
347                     close_with_forall vl (extract status amount vl lit),
348                            (* NCic.Implicit `Type, *)
349                     close_with_lambdas vl rewrite_step, body)
350     in 
351       aux false [] steps, proof_type
352   ;;
353