]> matita.cs.unibo.it Git - helm.git/blob - matita/components/ng_paramodulation/nCicProof.ml
Use of standard OCaml syntax
[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) in*)
202     let module Subst = FoSubst in
203     (*let compose subst1 subst2 =
204       let s1 = List.map (fun (x,t) -> (x, Subst.apply_subst subst2 t)) subst1 in
205       let s2 = List.filter (fun (m,_) -> not (List.mem_assoc m s1)) subst2
206       in s1 @ s2
207     in*)
208     let position i l = 
209       let rec aux = function
210        | [] -> assert false
211        | (j,_) :: _ when i = j -> 1
212        | _ :: tl -> 1 + aux tl
213       in
214         aux l
215     in
216     let vars_of i l = fst (List.assoc i l) in
217     let ty_of i l = snd (List.assoc i l) in
218     let close_with_lambdas vl t = 
219       List.fold_left 
220        (fun t i -> 
221           NCic.Lambda ("x"^string_of_int i, NCic.Implicit `Type, t))
222        t vl  
223     in
224     let close_with_forall vl t = 
225       List.fold_left 
226        (fun t i -> 
227           NCic.Prod ("x"^string_of_int i, NCic.Implicit `Type, t))
228        t vl  
229     in
230     let get_literal id =
231       let (_, lit, vl, proof),_,_ = Terms.get_from_bag id bag in
232       let lit =match lit with 
233           | Terms.Predicate t -> t (* assert false *) 
234           | Terms.Equation (l,r,ty,_) -> 
235               Terms.Node [ Terms.Leaf (eqP()); ty; l; r]
236       in
237         lit, vl, proof
238     in
239     let proof_type =
240       let lit,_,_ = get_literal mp in
241       let lit = Subst.apply_subst subst lit in
242         extract status 0 [] lit in
243     (* composition of all subst acting on goal *)
244     let res_subst = subst
245     in
246     let rec aux ongoal seen = function
247       | [] -> NCic.Rel 1
248       | id :: tl ->
249           let amount = List.length seen in
250           let lit,vl,proof = get_literal id in
251           if not ongoal && id = mp then
252             let lit = Subst.apply_subst subst lit in 
253             let eq_ty = extract status amount [] lit in
254             let refl = 
255               if demod then NCic.Implicit `Term 
256               else mk_refl eq_ty in
257               (* prerr_endline ("Reached m point, id=" ^ (string_of_int id));
258                 (NCic.LetIn ("clause_" ^ string_of_int id, eq_ty, refl,
259                 aux true ((id,([],lit))::seen) (id::tl))) *)
260               let res = NCicSubstitution.subst status
261                 ~avoid_beta_redexes:true ~no_implicit:false refl
262                 (aux true ((id,([],lit))::seen) (id::tl))
263               in res
264           else
265           match proof with
266           | Terms.Exact _ when tl=[] ->
267               (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id)); *)
268               aux ongoal seen tl
269           | Terms.Step _ when tl=[] -> assert false
270           | Terms.Exact ft ->
271               (*
272                prerr_endline ("Exact for " ^ (string_of_int id));
273                NCic.LetIn ("clause_" ^ string_of_int id, 
274                  close_with_forall vl (extract status amount vl lit),
275                  close_with_lambdas vl (extract status amount vl ft),
276                  aux ongoal 
277                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
278                *)
279                let res = NCicSubstitution.subst status
280                  ~avoid_beta_redexes:true ~no_implicit:false
281                  (close_with_lambdas vl (extract status amount vl ft))
282                  (aux ongoal 
283                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
284                in res 
285           | Terms.Step (_, id1, id2, dir, pos, subst) ->
286               (* prerr_endline (if ongoal then "on goal" else "not on goal");
287                  prerr_endline (Pp.pp_substitution subst); *)
288               (* let subst = if ongoal then res_subst else subst in *)
289               let id, id1,(lit,vl,_proof) =
290                 if ongoal then
291                   let lit,vl,proof = get_literal id1 in
292                   id1,id,(Subst.apply_subst res_subst lit,
293                           Subst.filter res_subst vl, proof)
294                 else id,id1,(lit,vl,proof) in
295               (* free variables remaining in the goal should not
296                  be abstracted: we do not want to prove a generalization *)
297               (* prerr_endline ("variables = " ^ String.concat ","
298                 (List.map string_of_int vl)); *)
299               let vl = if ongoal then [] else vl in 
300               let proof_of_id id = 
301                 let vars = List.rev (vars_of id seen) in
302                 let args = List.map (Subst.apply_subst subst) vars in
303                 let args = List.map (extract status amount vl) args in
304                 let rel_for_id = NCic.Rel (List.length vl + position id seen) in
305                   if args = [] then rel_for_id                    
306                   else NCic.Appl (rel_for_id::args)
307               in
308               let p_id1 = proof_of_id id1 in
309               let p_id2 = proof_of_id id2 in
310               let pred, hole_type, l, r = 
311                 let id1_ty = ty_of id1 seen in
312                 let id2_ty,l,r = 
313                   match ty_of id2 seen with
314                   | Terms.Node [ _; t; l; r ] -> 
315                       extract status amount vl (Subst.apply_subst subst t),
316                       extract status amount vl (Subst.apply_subst subst l),
317                       extract status amount vl (Subst.apply_subst subst r)
318                   | _ -> assert false
319                 in
320                 (* let _ = prerr_endline ("body = " ^(Pp.pp_foterm id1_ty)) 
321                    in *)
322                 mk_predicate status
323                   id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
324                 id2_ty,
325                 l,r
326               in
327               let rewrite_step =
328                 if (ongoal=true) = (dir=Terms.Left2Right) then
329                   NCic.Appl 
330                     [eq_ind_r(); hole_type; r; pred; p_id1; l; p_id2]
331                 else
332                   NCic.Appl 
333                     [ eq_ind(); hole_type; l; pred; p_id1; r; p_id2]
334               in
335               let body = aux ongoal 
336                 ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl 
337               in 
338               let occ =
339                NCicUntrusted.count_occurrences status [] 1 body in
340                 if occ <= 1 then
341                   NCicSubstitution.subst status
342                     ~avoid_beta_redexes:true ~no_implicit:false
343                     (close_with_lambdas vl rewrite_step) body
344                 else
345                   NCic.LetIn ("clause_" ^ string_of_int id, 
346                     close_with_forall vl (extract status amount vl lit),
347                            (* NCic.Implicit `Type, *)
348                     close_with_lambdas vl rewrite_step, body)
349     in 
350       aux false [] steps, proof_type
351   ;;
352