]> matita.cs.unibo.it Git - helm.git/blob - matita/components/ng_paramodulation/nCicProof.ml
HUGE COMMIT:
[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
68    let mk_predicate status hole_type amount ft p1 vl =
69     let rec aux t p = 
70       match p with
71       | [] -> NCic.Rel 1
72       | n::tl ->
73           match t with
74           | Terms.Leaf _ 
75           | Terms.Var _ -> 
76               let module NCicBlob = NCicBlob.NCicBlob(
77                         struct
78                           let metasenv = [] let subst = [] let context = []
79                         end)
80                           in
81               let module Pp = Pp.Pp(NCicBlob) in  
82                prerr_endline ("term: " ^ Pp.pp_foterm ft);
83                prerr_endline ("path: " ^ String.concat "," 
84                  (List.map string_of_int p1));
85                prerr_endline ("leading to: " ^ Pp.pp_foterm t);
86                assert false
87           | Terms.Node l -> 
88               let l = 
89                 HExtlib.list_mapi
90                   (fun t i -> 
91                     if i = n then aux t tl 
92                     else extract status amount (0::vl) t)
93                   l
94               in            
95               NCic.Appl l
96     in
97       NCic.Lambda("x", hole_type, aux ft (List.rev p1))
98     ;;
99
100   let dag = 
101    let uri = NUri.uri_of_string "cic:/matita/ng/sets/setoids/prop1.con" in
102    let ref = NReference.reference_of_spec uri (NReference.Fix(0,2,4)) in
103      NCic.Const ref
104   ;;
105
106   (*
107   let eq_setoid = 
108    let uri = NUri.uri_of_string "cic:/matita/ng/sets/setoids/eq.con" in
109    let ref = NReference.reference_of_spec uri (NReference.Fix(0,0,2)) in
110      NCic.Const ref
111   ;;
112   *)
113
114   let sym eq = 
115    let u= NUri.uri_of_string "cic:/matita/ng/properties/relations/sym.con" in
116    let u = NReference.reference_of_spec u (NReference.Fix(0,1,3)) in
117      NCic.Appl[NCic.Const u; NCic.Implicit `Type; NCic.Implicit `Term;
118      NCic.Implicit `Term; NCic.Implicit `Term; eq]; 
119   ;;
120
121   let eq_morphism1 eq = 
122    let u= NUri.uri_of_string "cic:/matita/ng/sets/setoids/eq_is_morphism1.con" in
123    let u = NReference.reference_of_spec u (NReference.Def 4) in
124      NCic.Appl[NCic.Const u; NCic.Implicit `Term; NCic.Implicit `Term;
125      NCic.Implicit `Term; NCic.Implicit `Term; eq]; 
126   ;;
127
128   let eq_morphism2 eq = 
129    let u= NUri.uri_of_string "cic:/matita/ng/sets/setoids/eq_is_morphism2.con" in
130    let u = NReference.reference_of_spec u (NReference.Def 4) in
131      NCic.Appl[NCic.Const u; NCic.Implicit `Term; NCic.Implicit `Term;
132      NCic.Implicit `Term; eq; NCic.Implicit `Term]; 
133   ;;
134
135   let trans eq p = 
136    let u= NUri.uri_of_string "cic:/matita/ng/properties/relations/trans.con" in
137    let u = NReference.reference_of_spec u (NReference.Fix(0,1,3)) in
138      NCic.Appl[NCic.Const u; NCic.Implicit `Type; NCic.Implicit `Term;
139      NCic.Implicit `Term; NCic.Implicit `Term; NCic.Implicit `Term; eq]
140   ;;
141
142   let iff1 eq p = 
143    let uri = NUri.uri_of_string "cic:/matita/ng/logic/connectives/if.con" in
144    let ref = NReference.reference_of_spec uri (NReference.Fix(0,2,1)) in
145      NCic.Appl[NCic.Const ref; NCic.Implicit `Type; NCic.Implicit `Type; 
146                eq; p]; 
147   ;;
148
149 (*
150   let mk_refl = function
151       | NCic.Appl [_; _; x; _] -> 
152    let uri= NUri.uri_of_string "cic:/matita/ng/properties/relations/refl.con" in
153    let ref = NReference.reference_of_spec uri (NReference.Fix(0,1,3)) in
154     NCic.Appl[NCic.Const ref; NCic.Implicit `Type; NCic.Implicit `Term;
155     NCic.Implicit `Term(*x*)]
156       | _ -> assert false
157 *)   
158
159   let mk_refl = function
160     | NCic.Appl [_; ty; l; _]
161       -> NCic.Appl [eq_refl();ty;l]
162     | _ -> assert false
163
164
165    let mk_morphism status eq amount ft pl vl =
166     let rec aux t p = 
167       match p with
168       | [] -> eq
169       | n::tl ->
170           prerr_endline (string_of_int n);
171           match t with
172           | Terms.Leaf _ 
173           | Terms.Var _ -> assert false
174           | Terms.Node [] -> assert false
175           | Terms.Node [ Terms.Leaf eqt ; _; l; r ]
176              when (eqP ()) = eqt ->
177                if n=2 then eq_morphism1 (aux l tl)
178                else eq_morphism2 (aux r tl)
179           | Terms.Node (f::l) ->
180              snd (
181               List.fold_left
182                (fun (i,acc) t ->
183                  i+1,
184                        let f = extract status amount vl f in
185                   if i = n then
186                    let imp = NCic.Implicit `Term in
187                     NCic.Appl (dag::imp::imp::imp(* f *)::imp::imp::
188                                [aux t tl])
189                   else
190                     NCicUntrusted.mk_appl acc [extract status amount vl t]
191                ) (1,extract status amount vl f) l)
192     in aux ft (List.rev pl)
193     ;;
194
195   let mk_proof status ?(demod=false) (bag : NCic.term Terms.bag) mp subst steps=
196     let module NCicBlob = 
197        NCicBlob.NCicBlob(
198          struct
199            let metasenv = [] let subst = [] let context = []
200          end)
201      in
202      let  module Pp = Pp.Pp(NCicBlob) 
203      in
204     let module Subst = FoSubst in
205     let position i l = 
206       let rec aux = function
207        | [] -> assert false
208        | (j,_) :: tl when i = j -> 1
209        | _ :: tl -> 1 + aux tl
210       in
211         aux l
212     in
213     let vars_of i l = fst (List.assoc i l) in
214     let ty_of i l = snd (List.assoc i l) in
215     let close_with_lambdas vl t = 
216       List.fold_left 
217        (fun t i -> 
218           NCic.Lambda ("x"^string_of_int i, NCic.Implicit `Type, t))
219        t vl  
220     in
221     let close_with_forall vl t = 
222       List.fold_left 
223        (fun t i -> 
224           NCic.Prod ("x"^string_of_int i, NCic.Implicit `Type, t))
225        t vl  
226     in
227     let get_literal id =
228       let (_, lit, vl, proof),_,_ = Terms.get_from_bag id bag in
229       let lit =match lit with 
230           | Terms.Predicate t -> t (* assert false *) 
231           | Terms.Equation (l,r,ty,_) -> 
232               Terms.Node [ Terms.Leaf eqP(); ty; l; r]
233       in
234         lit, vl, proof
235     in
236     let proof_type =
237       let lit,_,_ = get_literal mp in
238       let lit = Subst.apply_subst subst lit in
239         extract status 0 [] lit in
240     (* composition of all subst acting on goal *)
241     let res_subst =
242       let rec rsaux ongoal acc = 
243         function
244           | [] -> acc (* is the final subst for refl *)
245           | id::tl when ongoal ->
246             let lit,vl,proof = get_literal id in
247               (match proof with
248                 | Terms.Exact _ -> rsaux ongoal acc tl
249                 | Terms.Step (_, _, _, _, _, s) ->
250                     rsaux ongoal (s@acc) tl)
251           | id::tl -> 
252             (* subst is the the substitution for refl *)
253             rsaux (id=mp) subst tl
254       in 
255       let r = rsaux false [] steps in
256         (* prerr_endline ("res substitution: " ^ Pp.pp_substitution r);
257            prerr_endline ("\n" ^ "subst: " ^ Pp.pp_substitution subst); *)
258       r
259     in
260     let rec aux ongoal seen = function
261       | [] -> NCic.Rel 1
262       | id :: tl ->
263           let amount = List.length seen in
264           let lit,vl,proof = get_literal id in
265           if not ongoal && id = mp then
266             let lit = Subst.apply_subst subst lit in 
267             let eq_ty = extract status amount [] lit in
268             let refl = 
269               if demod then NCic.Implicit `Term 
270               else mk_refl eq_ty in
271              (* prerr_endline ("Reached m point, id=" ^ (string_of_int id));
272                 (NCic.LetIn ("clause_" ^ string_of_int id, eq_ty, refl,
273                 aux true ((id,([],lit))::seen) (id::tl))) *)
274               NCicSubstitution.subst status
275                 ~avoid_beta_redexes:true ~no_implicit:false refl
276                 (aux true ((id,([],lit))::seen) (id::tl))
277           else
278           match proof with
279           | Terms.Exact _ when tl=[] ->
280               (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id));*)
281               aux ongoal seen tl
282           | Terms.Step _ when tl=[] -> assert false
283           | Terms.Exact ft ->
284               (*
285                prerr_endline ("Exact for " ^ (string_of_int id));
286                NCic.LetIn ("clause_" ^ string_of_int id, 
287                  close_with_forall vl (extract status amount vl lit),
288                  close_with_lambdas vl (extract status amount vl ft),
289                  aux ongoal 
290                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
291               *)
292                NCicSubstitution.subst status
293                  ~avoid_beta_redexes:true ~no_implicit:false
294                  (close_with_lambdas vl (extract status amount vl ft))
295                  (aux ongoal 
296                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
297           | Terms.Step (_, id1, id2, dir, pos, subst) ->
298               let id, id1,(lit,vl,proof) =
299                 if ongoal then
300                   let lit,vl,proof = get_literal id1 in
301                   id1,id,(Subst.apply_subst res_subst lit, 
302                           Subst.filter res_subst vl, proof)
303                 else id,id1,(lit,vl,proof) in
304               (* free variables remaining in the goal should not
305                  be abstracted: we do not want to prove a generalization *)
306               let vl = if ongoal then [] else vl in 
307               let proof_of_id id = 
308                 let vars = List.rev (vars_of id seen) in
309                 let args = List.map (Subst.apply_subst subst) vars in
310                 let args = List.map (extract status amount vl) args in
311                 let rel_for_id = NCic.Rel (List.length vl + position id seen) in
312                   if args = [] then rel_for_id                    
313                   else NCic.Appl (rel_for_id::args)
314               in
315               let p_id1 = proof_of_id id1 in
316               let p_id2 = proof_of_id id2 in
317 (*
318               let morphism, l, r =
319                 let p =                 
320                   if (ongoal=true) = (dir=Terms.Left2Right) then
321                     p_id2 
322                   else sym p_id2 in
323                 let id1_ty = ty_of id1 seen in
324                 let id2_ty,l,r = 
325                   match ty_of id2 seen with
326                   | Terms.Node [ _; t; l; r ] -> 
327                       extract status amount vl (Subst.apply_subst subst t),
328                       extract status amount vl (Subst.apply_subst subst l),
329                       extract status amount vl (Subst.apply_subst subst r)
330                   | _ -> assert false
331                 in
332                   (*prerr_endline "mk_predicate :";
333                   if ongoal then prerr_endline "ongoal=true"
334                   else prerr_endline "ongoal=false";
335                   prerr_endline ("id=" ^ string_of_int id);
336                   prerr_endline ("id1=" ^ string_of_int id1);
337                   prerr_endline ("id2=" ^ string_of_int id2);
338                   prerr_endline ("Positions :" ^
339                                    (String.concat ", "
340                                       (List.map string_of_int pos)));*)
341                 mk_morphism
342                   p amount (Subst.apply_subst subst id1_ty) pos vl,
343                 l,r
344               in
345               let rewrite_step = iff1 morphism p_id1
346               in
347 *)
348               let pred, hole_type, l, r = 
349                 let id1_ty = ty_of id1 seen in
350                 let id2_ty,l,r = 
351                   match ty_of id2 seen with
352                   | Terms.Node [ _; t; l; r ] -> 
353                       extract status amount vl (Subst.apply_subst subst t),
354                       extract status amount vl (Subst.apply_subst subst l),
355                       extract status amount vl (Subst.apply_subst subst r)
356                   | _ -> assert false
357                 in
358                   (*
359                   prerr_endline "mk_predicate :";
360                   if ongoal then prerr_endline "ongoal=true"
361                   else prerr_endline "ongoal=false";
362                   prerr_endline ("id=" ^ string_of_int id);
363                   prerr_endline ("id1=" ^ string_of_int id1 
364                                  ^": " ^ Pp.pp_foterm id1_ty);
365                   prerr_endline ("id2=" ^ string_of_int id2
366                                  ^ ": " ^ NCicPp.ppterm [][][] id2_ty);
367                   prerr_endline ("Positions :" ^
368                                    (String.concat ", "
369                                       (List.map string_of_int pos)));*)
370                 mk_predicate status
371                   id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
372                 id2_ty,
373                 l,r
374               in
375               let rewrite_step =
376                 if (ongoal=true) = (dir=Terms.Left2Right) then
377                   NCic.Appl 
378                     [eq_ind_r(); hole_type; r; pred; p_id1; l; p_id2]
379                 else
380                   NCic.Appl 
381                     [ eq_ind(); hole_type; l; pred; p_id1; r; p_id2]
382               in
383               let body = aux ongoal 
384                 ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl 
385               in 
386               let occ =
387                NCicUntrusted.count_occurrences status [] 1 body in
388                 if occ <= 1 then
389                   NCicSubstitution.subst status
390                     ~avoid_beta_redexes:true ~no_implicit:false
391                     (close_with_lambdas vl rewrite_step) body
392                 else
393                   NCic.LetIn ("clause_" ^ string_of_int id, 
394                     close_with_forall vl (extract status amount vl lit),
395                            (* NCic.Implicit `Type, *)
396                     close_with_lambdas vl rewrite_step, body)
397     in 
398       aux false [] steps, proof_type
399   ;;
400