]> matita.cs.unibo.it Git - helm.git/commitdiff
- procedural: basic support for lapply (solves a problem in the reconstruction of...
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 23 Jun 2009 21:00:06 +0000 (21:00 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 23 Jun 2009 21:00:06 +0000 (21:00 +0000)
- lambda-delta: bugfix in line counts and nodes counts

helm/software/components/acic_procedural/procedural2.ml
helm/software/components/acic_procedural/proceduralTeX.ml
helm/software/components/acic_procedural/proceduralTypes.ml
helm/software/components/acic_procedural/proceduralTypes.mli
helm/software/lambda-delta/basic_ag/bagEnvironment.ml
helm/software/lambda-delta/toplevel/metaOutput.ml
helm/software/lambda-delta/toplevel/top.ml

index 07c1da5f091c312f4c9518d5105cdb5e65cb6156..ff8f864eabb31320e8af90737ae1501be4085434 100644 (file)
@@ -211,24 +211,28 @@ let are_convertible st pred sx dx =
 
 let anonymous_premise = C.Name "UNNAMED"
 
-let mk_exp_args hd tl classes synth qs =
+let mk_lapply_args hd tl classes = 
+   let map _ = Cn.meta "" in
+   let args = List.rev_map map tl in
+   if args = [] then hd else C.AAppl ("", hd :: args)
+
+let mk_apply_args hd tl classes synth qs =
    let exp = ref 0 in
-   let meta id = C.AImplicit (id, None) in
    let map v (cl, b) =
       if I.overlaps synth cl
-         then if b then v, v else meta "", v
-         else meta "", meta ""
+         then if b then v, v else Cn.meta "", v
+         else Cn.meta "", Cn.meta ""
    in
    let rec rev a = function
       | []       -> a
       | hd :: tl -> 
-         if snd hd <> meta "" then incr exp;
+         if snd hd <> Cn.meta "" then incr exp;
          rev (snd hd :: a) tl 
    in
    let rec aux = function
       | []       -> []
       | hd :: tl -> 
-         if fst hd = meta "" then aux tl else rev [] (hd :: tl)
+         if fst hd = Cn.meta "" then aux tl else rev [] (hd :: tl)
    in
    let args = T.list_rev_map2 map tl classes in
    let args = aux args in
@@ -359,6 +363,26 @@ and proc_letin st what name v w t =
                  mk_fwd_rewrite st dtext intro tl true v t ity ety
               | C.AAppl (_, hd :: tl) when is_fwd_rewrite_left st hd tl  ->
                  mk_fwd_rewrite st dtext intro tl false v t ity ety
+              | C.AAppl (_, hd :: tl)                                    ->
+                  let ty = match get_inner_types st hd with
+                     | Some (ity, _) -> H.cic ity 
+                    | None          -> get_type "TC3" st hd 
+                  in
+                 let classes, _ = Cl.classify st.context ty in
+                  let parsno, argsno = List.length classes, List.length tl in
+                  let decurry = parsno - argsno in
+                 if decurry <> 0 then begin              
+(* FG: we fall back in the cut case *)              
+                    assert (Ut.is_sober st.context (H.cic ety));
+                    let ety = H.acic_bc st.context ety in
+                    let qs = [proc_proof (next st) v; [T.Id ""]] in
+                    st, [T.Branch (qs, ""); T.Cut (intro, ety, dtext)]
+                 end else
+                 let names, synth = get_sub_names hd tl, I.S.empty in
+                 let qs = proc_bkd_proofs (next st) synth names classes tl in
+                  let hd = mk_lapply_args hd tl classes in
+                 let qs = [T.Id ""] :: qs in
+                 st, [T.Branch (qs, ""); T.LApply (intro, hd, dtext)]
               | v                                                        ->
                  assert (Ut.is_sober st.context (H.cic ety));
                  let ety = H.acic_bc st.context ety in
@@ -439,7 +463,7 @@ and proc_appl st what hd tl =
            in
            if List.length qs <> List.length names then
               let qs = proc_bkd_proofs (next st) synth [] classes tl in
-              let b, hd, qs = mk_exp_args hd tl classes synth qs in
+              let b, hd, qs = mk_apply_args hd tl classes synth qs in
               script @ [tactic b hd (dtext ^ text); T.Branch (qs, "")]
            else if is_rewrite_right st hd then 
               script2 @ mk_rewrite st dtext where qs tl2 false what ity
@@ -454,7 +478,7 @@ and proc_appl st what hd tl =
         | _                                       ->
            let names = get_sub_names hd tl in
            let qs = proc_bkd_proofs (next st) synth names classes tl in
-           let b, hd, qs = mk_exp_args hd tl classes synth qs in
+           let b, hd, qs = mk_apply_args hd tl classes synth qs in
            script @ [tactic b hd (dtext ^ text); T.Branch (qs, "")]
    else
       [T.Exact (what, dtext)]
index 5f4d8d150d0cff31f3e2dbe3018d6553412b186e..294fefb045f747da13590401c42e01be32a5655a 100644 (file)
@@ -202,6 +202,8 @@ let rec xl frm = function
       F.fprintf frm "\\Intro{%a}{%a}" xx r xl l
    | T.LetIn (r, v, _) :: l                                ->
       F.fprintf frm "\\Pose{%a}{%a}{%a}" xx r xat v xl l
+   | T.LApply (r, v, _) :: l                               ->
+      F.fprintf frm "\\LApply{%a}{%a}{%a}" xx r xat v xl l
    | T.Change (u, _, None, _, _) :: l                      ->
       F.fprintf frm "\\Change{%a}{}{%a}" xat u xl l
    | T.Change (u, _, Some (s, _), _, _) :: l               ->
index 8cf705d0a1c098ca87375397799d1997a147b400..c465315d4bc183b1e1a7c1ef8f852bff310dd829 100644 (file)
@@ -80,6 +80,7 @@ type step = Note of note
          | Intros of count option * name list * note
          | Cut of name * what * note
          | LetIn of name * what * note
+         | LApply of name * what * note
          | Rewrite of how * what * where * pattern * note
          | Elim of what * using option * pattern * note
          | Cases of what * pattern * note
@@ -207,6 +208,10 @@ let mk_letin name what punctation =
    let tactic = G.LetIn (floc, what, name) in
    mk_tactic tactic punctation
 
+let mk_lapply name what punctation =
+   let tactic = G.LApply (floc, false, None, [], what, name) in
+   mk_tactic tactic punctation
+
 let mk_rewrite direction what where pattern punctation =
    let direction = if direction then `RightToLeft else `LeftToRight in 
    let pattern, rename = match where with
@@ -276,6 +281,7 @@ let rec render_step sep a = function
    | Intros (c, ns, s)         -> mk_intros c ns sep :: mk_tacnote s a
    | Cut (n, t, s)             -> mk_cut n t sep :: mk_tacnote s a
    | LetIn (n, t, s)           -> mk_letin n t sep :: mk_tacnote s a
+   | LApply (n, t, s)          -> mk_lapply n t sep :: mk_tacnote s a
    | Rewrite (b, t, w, e, s)   -> mk_rewrite b t w e sep :: mk_tacnote s a
    | Elim (t, xu, e, s)        -> mk_elim t xu e sep :: mk_tacnote s a
    | Cases (t, e, s)           -> mk_cases t e sep :: mk_tacnote s a
@@ -344,6 +350,7 @@ let rec count_node a = function
    | Exact (t, _) 
    | Cut (_, t, _)
    | LetIn (_, t, _)
+   | LApply (_, t, _)
    | Apply (t, _)            -> count a (H.cic t)
    | Rewrite (_, t, _, p, _)
    | Elim (t, _, p, _)
@@ -366,6 +373,7 @@ let rec note_of_step = function
    | Intros (_, _, s)
    | Cut (_, _, s)
    | LetIn (_, _, s)
+   | LApply (_, _, s)
    | Rewrite (_, _, _, _, s)
    | Elim (_, _, _, s)
    | Cases (_, _, s)
index a1d28e966c2c8428116fbd2be067bb808da101bc..969492a627bd655b63adea098fcdd020cfccf202 100644 (file)
@@ -59,6 +59,7 @@ type step = Note of note
          | Intros of count option * name list * note
          | Cut of name * what * note
          | LetIn of name * what * note
+         | LApply of name * what * note
          | Rewrite of how * what * where * pattern * note
          | Elim of what * using option * pattern * note
          | Cases of what * pattern * note
index 5e6c9f1a0d58a7ce75df7c1e38f9cc1eda83f18d..f145661723be6b4664d4d841a82d38dd216c3a81 100644 (file)
@@ -18,7 +18,7 @@ exception ObjectNotFound of B.message
 
 let hsize = 7000 
 let env = H.create hsize
-let entry = ref 0
+let entry = ref 1
 
 (* Internal functions *******************************************************)
 
index 7349fa573d9cbca4ee41e456c2c17304d361dbef..4cec084b69c1241e544e151799aaad3ef4cecd38 100644 (file)
@@ -60,9 +60,7 @@ let rec count_term f c = function
       let f c = count_term f c t in
       count_term f c w
 
-let count_par f c (_, w) = 
-   let c = {c with nodes = succ c.nodes} in
-   count_term f c w
+let count_par f c (_, w) = count_term f c w
 
 let count_entry f c = function
    | _, pars, u, w, None        ->
@@ -84,7 +82,7 @@ let count_item f c = function
    | None   -> f c
 
 let print_counters f c =
-   let terms = c.tsorts + c.tgrefs + c.tgrefs + c.tappls + c.tabsts in
+   let terms = c.tsorts + c.tlrefs + c.tgrefs + c.tappls + c.tabsts in
    let pars = c.pabsts + c.pappls in
    let items = c.eabsts + c.eabbrs in
    let nodes = c.nodes + c.xnodes in
index 22bd5cbc7ef4ca103f35c4b2704a9fd3d73da274..1774cc991fe297bbf5c185ed0471b6bbeb671600 100644 (file)
@@ -87,7 +87,7 @@ try
            let f st _ = function
               | None           -> st
               | Some (i, u, _) -> 
-(*               Log.warn (P.sprintf "[%u] %s" i (U.string_of_uri u)); *)
+                 Log.warn (P.sprintf "[%u] %s" i (U.string_of_uri u)); 
                  st
            in
 (* stage 2 *)