]> matita.cs.unibo.it Git - helm.git/blob - helm/software/components/ng_paramodulation/nCicProof.ml
Various architectural changes
[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   let eqP () = 
15     let r = 
16       OCic2NCic.reference_of_oxuri 
17        (UriManager.uri_of_string 
18          "cic:/matita/logic/equality/eq.ind#xpointer(1/1)")
19     in
20     NCic.Const r
21   ;;
22
23   let eq_ind () = 
24     let r = 
25       OCic2NCic.reference_of_oxuri 
26        (UriManager.uri_of_string 
27          "cic:/matita/logic/equality/eq_ind.con")
28     in
29     NCic.Const r
30   ;;
31
32   let eq_ind_r () = 
33     let r = 
34       OCic2NCic.reference_of_oxuri 
35        (UriManager.uri_of_string 
36          "cic:/matita/logic/equality/eq_elim_r.con")
37     in
38     NCic.Const r
39   ;;
40
41   let eq_refl () = 
42     let r = 
43       OCic2NCic.reference_of_oxuri 
44        (UriManager.uri_of_string 
45          "cic:/matita/logic/equality/eq.ind#xpointer(1/1/1)")
46     in
47     NCic.Const r
48   ;;
49
50   let extract lift vl t =
51     let rec pos i = function 
52       | [] -> raise Not_found 
53       | j :: tl when j <> i -> 1+ pos i tl
54       | _ -> 1
55     in
56     let vl_len = List.length vl in
57     let rec extract = function
58       | Terms.Leaf x -> NCicSubstitution.lift (vl_len+lift) x
59       | Terms.Var j -> 
60            (try NCic.Rel (pos j vl) with Not_found -> NCic.Implicit `Term) 
61       | Terms.Node l -> NCic.Appl (List.map extract l)
62     in
63       extract t
64   ;;
65
66    let mk_predicate hole_type amount ft p1 vl =
67     let rec aux t p = 
68       match p with
69       | [] -> NCic.Rel 1
70       | n::tl ->
71           match t with
72           | Terms.Leaf _ 
73           | Terms.Var _ -> 
74               let module Pp = 
75                 Pp.Pp(NCicBlob.NCicBlob(
76                         struct
77                           let metasenv = [] let subst = [] let context = []
78                         end))
79               in  
80                prerr_endline ("term: " ^ Pp.pp_foterm ft);
81                prerr_endline ("path: " ^ String.concat "," 
82                  (List.map string_of_int p1));
83                prerr_endline ("leading to: " ^ Pp.pp_foterm t);
84                assert false
85           | Terms.Node l -> 
86               let l = 
87                 HExtlib.list_mapi
88                   (fun t i -> 
89                     if i = n then aux t tl 
90                     else extract amount (0::vl) t)
91                   l
92               in            
93               NCic.Appl l
94     in
95       NCic.Lambda("x", hole_type, aux ft (List.rev p1))
96     ;;
97
98   let mk_proof (bag : NCic.term Terms.bag) mp steps =
99     let module Subst = FoSubst in
100     let position i l = 
101       let rec aux = function
102        | [] -> assert false
103        | (j,_) :: tl when i = j -> 1
104        | _ :: tl -> 1 + aux tl
105       in
106         aux l
107     in
108     let vars_of i l = fst (List.assoc i l) in
109     let ty_of i l = snd (List.assoc i l) in
110     let close_with_lambdas vl t = 
111       List.fold_left 
112        (fun t i -> 
113           NCic.Lambda ("x"^string_of_int i, NCic.Implicit `Type, t))
114        t vl  
115     in
116     let close_with_forall vl t = 
117       List.fold_left 
118        (fun t i -> 
119           NCic.Prod ("x"^string_of_int i, NCic.Implicit `Type, t))
120        t vl  
121     in
122     let get_literal id =
123       let (_, nlit, plit, vl, proof),_,_ = Terms.get_from_bag id bag in
124       let lit = match nlit,plit with 
125          | [Terms.Equation (l,r,ty,_),_],[]
126          | [],[Terms.Equation (l,r,ty,_),_] ->
127               Terms.Node [ Terms.Leaf eqP(); ty; l; r]
128           | _ -> assert false
129       in
130         lit, vl, proof
131     in
132     let rec aux ongoal seen = function
133       | [] -> NCic.Rel 1
134       | id :: tl ->
135           let amount = List.length seen in
136           let lit,vl,proof = get_literal id in
137           if not ongoal && id = mp then
138             ((*prerr_endline ("Reached m point, id=" ^ (string_of_int id));*)
139              NCic.LetIn ("clause_" ^ string_of_int id, 
140                 extract amount [] lit, 
141                 (NCic.Appl [eq_refl();NCic.Implicit `Type;NCic.Implicit `Term]),
142                 aux true ((id,([],lit))::seen) (id::tl)))
143           else
144           match proof with
145           | Terms.Exact _ when tl=[] ->
146               (* prerr_endline ("Exact (tl=[]) for " ^ (string_of_int id));*)
147               aux ongoal seen tl
148           | Terms.Step _ when tl=[] -> assert false
149           | Terms.Exact ft ->
150               (* prerr_endline ("Exact for " ^ (string_of_int id));*)
151                NCic.LetIn ("clause_" ^ string_of_int id, 
152                  close_with_forall vl (extract amount vl lit),
153                  close_with_lambdas vl (extract amount vl ft),
154                  aux ongoal 
155                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
156           | Terms.Step (_, id1, id2, dir, pos, subst) ->
157               let id, id1,(lit,vl,proof) =
158                 if ongoal then id1,id,get_literal id1
159                 else id,id1,(lit,vl,proof)
160               in
161               let vl = if ongoal then [](*Subst.filter subst vl*) else vl in
162               let proof_of_id id = 
163                 let vars = List.rev (vars_of id seen) in
164                 let args = List.map (Subst.apply_subst subst) vars in
165                 let args = List.map (extract amount vl) args in
166                 let rel_for_id = NCic.Rel (List.length vl + position id seen) in
167                   if args = [] then rel_for_id              
168                   else NCic.Appl (rel_for_id::args)
169               in
170               let p_id1 = proof_of_id id1 in
171               let p_id2 = proof_of_id id2 in
172               let pred, hole_type, l, r = 
173                 let id1_ty = ty_of id1 seen in
174                 let id2_ty,l,r = 
175                   match ty_of id2 seen with
176                   | Terms.Node [ _; t; l; r ] -> 
177                       extract amount vl (Subst.apply_subst subst t),
178                       extract amount vl (Subst.apply_subst subst l),
179                       extract amount vl (Subst.apply_subst subst r)
180                   | _ -> assert false
181                 in
182                   (*prerr_endline "mk_predicate :";
183                   if ongoal then prerr_endline "ongoal=true"
184                   else prerr_endline "ongoal=false";
185                   prerr_endline ("id=" ^ string_of_int id);
186                   prerr_endline ("id1=" ^ string_of_int id1);
187                   prerr_endline ("id2=" ^ string_of_int id2);
188                   prerr_endline ("Positions :" ^
189                                    (String.concat ", "
190                                       (List.map string_of_int pos)));*)
191                 mk_predicate 
192                   id2_ty amount (Subst.apply_subst subst id1_ty) pos vl,
193                 id2_ty,
194                 l,r
195               in
196               let l, r, eq_ind = 
197                 if (ongoal=true) = (dir=Terms.Left2Right) then
198                   r,l,eq_ind_r ()
199                 else
200                   l,r,eq_ind ()
201               in
202                NCic.LetIn ("clause_" ^ string_of_int id, 
203                  close_with_forall vl (extract amount vl lit),
204                            (* NCic.Implicit `Type, *)
205                  close_with_lambdas vl 
206                    (NCic.Appl [ eq_ind ; hole_type; l; pred; p_id1; r; p_id2 ]),
207                  aux ongoal 
208                    ((id,(List.map (fun x -> Terms.Var x) vl,lit))::seen) tl)
209     in 
210       aux false [] steps
211   ;;
212