]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/components/tactics/primitiveTactics.ml
oblivion ugraph everywhere outside the kernel
[helm.git] / helm / software / components / tactics / primitiveTactics.ml
index c5aab0f5ff45eeebb63f608195552302d799df26..53bfc39a9cf0e236a0a1bc1d5159ea0fb105571e 100644 (file)
@@ -62,11 +62,11 @@ let lambda_abstract ?(howmany=(-1)) metasenv context newmeta ty mk_fresh_name =
            collect_context ctx (howmany - 1) do_whd t 
           in
            (context',ty,C.Lambda(n',s,bo))
-      | C.LetIn (n,s,t) ->
+      | C.LetIn (n,s,sty,t) ->
          let (context',ty,bo) =
-          collect_context ((Some (n,(C.Def (s,None))))::context) (howmany - 1) do_whd t
+          collect_context ((Some (n,(C.Def (s,sty))))::context) (howmany - 1) do_whd t
          in
-          (context',ty,C.LetIn(n,s,bo))
+          (context',ty,C.LetIn(n,s,sty,bo))
       | _ as t ->
         if howmany <= 0 then
          let irl =
@@ -102,7 +102,7 @@ let eta_expand metasenv context t arg =
     | C.Cast (te,ty) -> C.Cast (aux n te, aux n ty)
     | C.Prod (nn,s,t) -> C.Prod (nn, aux n s, aux (n+1) t)
     | C.Lambda (nn,s,t) -> C.Lambda (nn, aux n s, aux (n+1) t)
-    | C.LetIn (nn,s,t) -> C.LetIn (nn, aux n s, aux (n+1) t)
+    | C.LetIn (nn,s,ty,t) -> C.LetIn (nn, aux n s, aux n ty, aux (n+1) t)
     | C.Appl l -> C.Appl (List.map (aux n) l)
     | C.Const (uri,exp_named_subst) ->
        let exp_named_subst' = aux_exp_named_subst n exp_named_subst in
@@ -136,7 +136,7 @@ let eta_expand metasenv context t arg =
    List.map (function uri,t -> uri,aux n t)
   in
    let argty,_ = 
-    T.type_of_aux' metasenv context arg CicUniv.empty_ugraph (* TASSI: FIXME *)
+    T.type_of_aux' metasenv context arg CicUniv.oblivion_ugraph (* TASSI: FIXME *)
    in
     let fresh_name =
      FreshNamesGenerator.mk_fresh_name ~subst:[]
@@ -159,15 +159,13 @@ let classify_metas newmeta in_subst_domain subst_in metasenv =
            match entry with
               Some (n,Cic.Decl s) ->
                Some (n,Cic.Decl (subst_in canonical_context' s))
-            | Some (n,Cic.Def (s,None)) ->
-               Some (n,Cic.Def ((subst_in canonical_context' s),None))
             | None -> None
-            | Some (n,Cic.Def (bo,Some ty)) ->
+            | Some (n,Cic.Def (bo,ty)) ->
                Some
                 (n,
                   Cic.Def
                    (subst_in canonical_context' bo,
-                    Some (subst_in canonical_context' ty)))
+                    subst_in canonical_context' ty))
           in
            entry'::canonical_context'
         ) canonical_context []
@@ -184,7 +182,7 @@ let
 =
  let module C = Cic in
   let params =
-    let o,_ = CicEnvironment.get_obj CicUniv.empty_ugraph uri in
+    let o,_ = CicEnvironment.get_obj CicUniv.oblivion_ugraph uri in
     CicUtil.params_of_obj o
   in
    let exp_named_subst_diff,new_fresh_meta,newmetasenvfragment,exp_named_subst'=
@@ -196,7 +194,7 @@ let
          [],[] -> []
        | uri::tl,[] ->
           let ty =
-            let o,_ = CicEnvironment.get_obj CicUniv.empty_ugraph uri in
+            let o,_ = CicEnvironment.get_obj CicUniv.oblivion_ugraph uri in
               match o with
                   C.Variable (_,_,ty,_,_) ->
                     CicSubstitution.subst_vars !exp_named_subst_diff ty
@@ -244,7 +242,7 @@ let new_metasenv_and_unify_and_t newmeta' metasenv' subst context term' ty termt
     goal_arity in
   let subst,newmetasenv',_ = 
    CicUnification.fo_unif_subst 
-     subst context newmetasenv consthead ty CicUniv.empty_ugraph
+     subst context newmetasenv consthead ty CicUniv.oblivion_ugraph
   in
   let t = 
     if List.length arguments = 0 then term' else Cic.Appl (term'::arguments)
@@ -261,7 +259,7 @@ let apply_with_subst ~term ~subst ~maxmeta (proof, goal) =
  let module T = CicTypeChecker in
  let module R = CicReduction in
  let module C = Cic in
-  let (_,metasenv,_,_, _) = proof in
+  let (_,metasenv,_subst,_,_, _) = proof in
   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
   let newmeta = max (CicMkImplicit.new_meta metasenv subst) maxmeta in
    let exp_named_subst_diff,newmeta',newmetasenvfragment,term' =
@@ -298,7 +296,7 @@ let apply_with_subst ~term ~subst ~maxmeta (proof, goal) =
    in
    let metasenv' = metasenv@newmetasenvfragment in
    let termty,_ = 
-     CicTypeChecker.type_of_aux' metasenv' context term' CicUniv.empty_ugraph
+     CicTypeChecker.type_of_aux' metasenv' context term' CicUniv.oblivion_ugraph
    in
    let termty =
      CicSubstitution.subst_vars exp_named_subst_diff termty in
@@ -325,7 +323,7 @@ let apply_with_subst ~term ~subst ~maxmeta (proof, goal) =
    let subst_in =
      (* if we just apply the subtitution, the type is irrelevant:
               we may use Implicit, since it will be dropped *)
-     CicMetaSubst.apply_subst ((metano,(context,bo',Cic.Implicit None))::subst)
+      ((metano,(context,bo',Cic.Implicit None))::subst)
    in
    let (newproof, newmetasenv''') = 
     ProofEngineHelpers.subst_meta_and_metasenv_in_proof proof metano subst_in
@@ -373,7 +371,7 @@ let intros_tac ?howmany ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_
  =
   let module C = Cic in
   let module R = CicReduction in
-   let (_,metasenv,_,_, _) = proof in
+   let (_,metasenv,_subst,_,_, _) = proof in
    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
     let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
      let (context',ty',bo') =
@@ -393,7 +391,7 @@ let cut_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:
   term (proof, goal)
  =
   let module C = Cic in
-   let curi,metasenv,pbo,pty, attrs = proof in
+   let curi,metasenv,_subst,pbo,pty, attrs = proof in
    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
     let newmeta1 = ProofEngineHelpers.new_meta_of_proof ~proof in
     let newmeta2 = newmeta1 + 1 in
@@ -409,11 +407,9 @@ let cut_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:
       CicMkImplicit.identity_relocation_list_for_metavariable context
     in
      let newmeta1ty = CicSubstitution.lift 1 ty in
-     let bo' =
-      C.Appl
-       [C.Lambda (fresh_name,term,C.Meta (newmeta1,irl1)) ;
-        C.Meta (newmeta2,irl2)]
-     in
+      let bo' = 
+        Cic.LetIn (fresh_name, C.Meta (newmeta2,irl2), term, C.Meta (newmeta1,irl1))
+      in
       let (newproof, _) =
        ProofEngineHelpers.subst_meta_in_proof proof metano bo'
         [newmeta2,context,term; newmeta1,context_for_newmeta1,newmeta1ty];
@@ -428,7 +424,7 @@ let letin_tac ?(mk_fresh_name_callback=FreshNamesGenerator.mk_fresh_name ~subst:
   term (proof, goal)
  =
   let module C = Cic in
-   let curi,metasenv,pbo,pty, attrs = proof in
+   let curi,metasenv,_subst,pbo,pty, attrs = proof in
    (* occur check *)
    let occur i t =
      let m = CicUtil.metas_of_term t in 
@@ -439,19 +435,19 @@ let letin_tac ?(mk_fresh_name_callback=FreshNamesGenerator.mk_fresh_name ~subst:
      raise 
        (ProofEngineTypes.Fail (lazy
          "You can't letin a term containing the current goal"));
-    let _,_ =
-      CicTypeChecker.type_of_aux' metasenv context term CicUniv.empty_ugraph in
+    let tty,_ =
+      CicTypeChecker.type_of_aux' metasenv context term CicUniv.oblivion_ugraph in
      let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
      let fresh_name =
       mk_fresh_name_callback metasenv context (Cic.Name "Hletin") ~typ:term in
      let context_for_newmeta =
-      (Some (fresh_name,C.Def (term,None)))::context in
+      (Some (fresh_name,C.Def (term,tty)))::context in
      let irl =
       CicMkImplicit.identity_relocation_list_for_metavariable
        context_for_newmeta
      in
       let newmetaty = CicSubstitution.lift 1 ty in
-      let bo' = C.LetIn (fresh_name,term,C.Meta (newmeta,irl)) in
+      let bo' = C.LetIn (fresh_name,term,tty,C.Meta (newmeta,irl)) in
        let (newproof, _) =
          ProofEngineHelpers.subst_meta_in_proof
            proof metano bo'[newmeta,context_for_newmeta,newmetaty]
@@ -464,11 +460,11 @@ let letin_tac ?(mk_fresh_name_callback=FreshNamesGenerator.mk_fresh_name ~subst:
 let exact_tac ~term =
  let exact_tac ~term (proof, goal) =
   (* Assumption: the term bo must be closed in the current context *)
-  let (_,metasenv,_,_, _) = proof in
+  let (_,metasenv,_subst,_,_, _) = proof in
   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
   let module T = CicTypeChecker in
   let module R = CicReduction in
-  let ty_term,u = T.type_of_aux' metasenv context term CicUniv.empty_ugraph in
+  let ty_term,u = T.type_of_aux' metasenv context term CicUniv.oblivion_ugraph in
   let b,_ = R.are_convertible context ty_term ty u in (* TASSI: FIXME *)
   if b then
    begin
@@ -494,10 +490,65 @@ module S   = CicSubstitution
 module T   = Tacticals
 module RT  = ReductionTactics
 
+let rec args_init n f =
+   if n <= 0 then [] else f n :: args_init (pred n) f
+
+let mk_predicate_for_elim 
+ ~context ~metasenv ~ugraph ~goal ~arg ~using ~cpattern ~args_no = 
+   let instantiated_eliminator =
+      let f n = if n = 1 then arg else C.Implicit None in
+      C.Appl (using :: args_init args_no f)
+   in
+   let _actual_arg, iety, _metasenv', _ugraph = 
+      CicRefine.type_of_aux' metasenv context instantiated_eliminator ugraph
+   in
+   let _actual_meta, actual_args = match iety with
+      | C.Meta (i, _)                  -> i, []
+      | C.Appl (C.Meta (i, _) :: args) -> i, args
+      | _                              -> assert false
+   in
+(* let _, upto = PEH.split_with_whd (List.nth splits pred_pos) in *)
+   let rec mk_pred metasenv context' pred arg' cpattern' = function
+      | []           -> metasenv, pred, arg'
+      | arg :: tail -> 
+(* FG: we find the predicate for the eliminator as in the rewrite tactic ****)
+        let argty, _ugraph = TC.type_of_aux' metasenv context arg ugraph in
+         let argty = CicReduction.whd context argty in         
+              let fresh_name = 
+            FreshNamesGenerator.mk_fresh_name 
+            ~subst:[] metasenv context' C.Anonymous ~typ:argty
+         in
+        let hyp = Some (fresh_name, C.Decl argty) in
+         let lazy_term c m u =  
+            let distance  = List.length c - List.length context in
+            S.lift distance arg, m, u
+         in
+         let pattern = Some lazy_term, [], Some cpattern' in
+         let subst, metasenv, _ugraph, _conjecture, selected_terms =
+            ProofEngineHelpers.select
+            ~metasenv ~ugraph ~conjecture:(0, context, pred) ~pattern
+         in
+         let metasenv = MS.apply_subst_metasenv subst metasenv in  
+         let map (_context_of_t, t) l = t :: l in
+         let what = List.fold_right map selected_terms [] in
+         let arg' = MS.apply_subst subst arg' in
+         let argty = MS.apply_subst subst argty in
+         let pred = PER.replace_with_rel_1_from ~equality:(==) ~what 1 pred in
+         let pred = MS.apply_subst subst pred in
+         let pred = C.Lambda (fresh_name, argty, pred) in
+        let cpattern' = C.Lambda (C.Anonymous, C.Implicit None, cpattern') in
+         mk_pred metasenv (hyp :: context') pred arg' cpattern' tail 
+   in
+   let metasenv, pred, arg = 
+      mk_pred metasenv context goal arg cpattern (List.rev actual_args)
+   in
+   HLog.debug ("PREDICATE: " ^ CicPp.ppterm ~metasenv pred ^ " ARGS: " ^ String.concat " " (List.map (CicPp.ppterm ~metasenv) actual_args));
+   metasenv, pred, arg, actual_args
+
 let beta_after_elim_tac upto predicate =
    let beta_after_elim_tac status =
       let proof, goal = status in
-      let _, metasenv, _, _, _ = proof in
+      let _, metasenv, _subst, _, _, _ = proof in
       let _, _, ty = CicUtil.lookup_meta goal metasenv in
       let mk_pattern ~equality ~upto ~predicate ty =
          (* code adapted from ProceduralConversion.generalize *)
@@ -542,9 +593,9 @@ let beta_after_elim_tac upto predicate =
             | C.Lambda (_, s, t) ->
                let s, t = gen_term k s, gen_term (succ k) t in
                if is_meta [s; t] then meta else C.Lambda (anon, s, t)
-            | C.LetIn (_, s, t) -> 
-               let s, t = gen_term k s, gen_term (succ k) t in
-               if is_meta [s; t] then meta else C.LetIn (anon, s, t)
+            | C.LetIn (_, s, ty, t) -> 
+               let s,ty,t = gen_term k s, gen_term k ty, gen_term (succ k) t in
+               if is_meta [s; t] then meta else C.LetIn (anon, s, ty, t)
             | C.Fix (i, fl) -> C.Fix (i, List.map (gen_fix (List.length fl) k) fl)
             | C.CoFix (i, fl) -> C.CoFix (i, List.map (gen_cofix (List.length fl) k) fl)
          in
@@ -559,12 +610,12 @@ let beta_after_elim_tac upto predicate =
    
 let elim_tac ?using ?(pattern = PET.conclusion_pattern None) term = 
  let elim_tac (proof, goal) =
-   let cpatt = match pattern with 
-      | None, [], Some cpatt -> cpatt
-      | _                    -> raise (PET.Fail (lazy "not implemented"))
+   let cpattern = match pattern with 
+      | None, [], Some cpattern -> cpattern
+      | _                       -> raise (PET.Fail (lazy "not implemented"))
    in    
-   let ugraph = CicUniv.empty_ugraph in
-   let curi, metasenv, proofbo, proofty, attrs = proof in
+   let ugraph = CicUniv.oblivion_ugraph in
+   let curi, metasenv, _subst, proofbo, proofty, attrs = proof in
    let conjecture = CicUtil.lookup_meta goal metasenv in
    let metano, context, ty = conjecture in 
     let termty,_ugraph = TC.type_of_aux' metasenv context term ugraph in
@@ -610,67 +661,20 @@ let elim_tac ?using ?(pattern = PET.conclusion_pattern None) term =
          TC.type_of_aux' metasenv' context eliminator_ref ugraph in
 (* FG: ADDED PART ***********************************************************)
 (* FG: we can not assume eliminator is the default eliminator ***************)
-   let rec args_init n f =
-      if n <= 0 then [] else f n :: args_init (pred n) f
-   in
    let splits, args_no = PEH.split_with_whd (context, ety) in
    let pred_pos = match List.hd splits with
       | _, C.Rel i when i > 1 && i <= args_no -> i
       | _, C.Appl (C.Rel i :: _) when i > 1 && i <= args_no -> i
       | _ -> raise NotAnEliminator
    in
-   let upto, metasenv', pred, term = match pattern with 
+   let metasenv', pred, term, actual_args = match pattern with 
       | None, [], Some (C.Implicit (Some `Hole)) ->
-         0, metasenv', C.Implicit None, term
+         metasenv', C.Implicit None, term, []
       | _                                        ->
-         let instantiated_eliminator =
-            let f n = if n = 1 then term else C.Implicit None in
-            C.Appl (eliminator_ref :: args_init args_no f)
-         in
-         let _actual_term, iety, _metasenv'', _ugraph = 
-            CicRefine.type_of_aux' metasenv' context instantiated_eliminator ugraph
-         in
-         let _actual_meta, actual_args = match iety with
-            | C.Meta (i, _)                  -> i, []
-            | C.Appl (C.Meta (i, _) :: args) -> i, args
-            | _                              -> assert false
-         in
-         (* let _, upto = PEH.split_with_whd (List.nth splits pred_pos) in *)
-         let upto = List.length actual_args in
-         let rec mk_pred metasenv context' pred term' = function
-            | []           -> metasenv, pred, term'
-            | term :: tail -> 
-(* FG: we find the predicate for the eliminator as in the rewrite tactic ****)
-               let termty, _ugraph = TC.type_of_aux' metasenv context' term ugraph in
-               let termty = CicReduction.whd context' termty in         
-              let fresh_name = 
-                  FreshNamesGenerator.mk_fresh_name 
-                  ~subst:[] metasenv context' C.Anonymous ~typ:termty
-               in
-              let hyp = Some (fresh_name, C.Decl termty) in
-               let lazy_term c m u =  
-                  let distance  = List.length c - List.length context in
-                  S.lift distance term, m, u
-               in
-               let pattern = Some lazy_term, [], Some cpatt in
-               let subst, metasenv, _ugraph, _conjecture, selected_terms =
-                  ProofEngineHelpers.select
-                  ~metasenv ~ugraph ~conjecture:(metano, context, pred) ~pattern
-               in
-               let metasenv = MS.apply_subst_metasenv subst metasenv in  
-               let map (_context_of_t, t) l = t :: l in
-               let what = List.fold_right map selected_terms [] in
-               let term' = MS.apply_subst subst term' in
-               let termty = MS.apply_subst subst termty in
-               let pred = PER.replace_with_rel_1_from ~equality:(==) ~what 1 pred in
-               let pred = MS.apply_subst subst pred in
-               let pred = C.Lambda (fresh_name, termty, pred) in
-               mk_pred metasenv (hyp :: context') pred term' tail 
-         in
-         let metasenv', pred, term = mk_pred metasenv' context ty term actual_args in
-         HLog.debug ("PRED: " ^ CicPp.ppterm ~metasenv:metasenv' pred ^ " ARGS: " ^ String.concat " " (List.map (CicPp.ppterm ~metasenv:metasenv') actual_args));
-         upto, metasenv', pred, term
-      in
+         mk_predicate_for_elim 
+           ~args_no ~context ~ugraph ~cpattern
+           ~metasenv:metasenv' ~arg:term ~using:eliminator_ref ~goal:ty
+   in
 (* FG: END OF ADDED PART ****************************************************)
       let term_to_refine =
          let f n =
@@ -686,18 +690,19 @@ let elim_tac ?using ?(pattern = PET.conclusion_pattern None) term =
          ProofEngineHelpers.compare_metasenvs
             ~oldmetasenv:metasenv ~newmetasenv:metasenv''
       in
-      let proof' = curi,metasenv'',proofbo,proofty, attrs in
+      let proof' = curi,metasenv'',_subst,proofbo,proofty, attrs in
       let proof'', new_goals' =
          PET.apply_tactic (apply_tac ~term:refined_term) (proof',goal)
       in
       (* The apply_tactic can have closed some of the new_goals *)
       let patched_new_goals =
-         let (_,metasenv''',_,_, _) = proof'' in
+         let (_,metasenv''',_subst,_,_, _) = proof'' in
          List.filter
             (function i -> List.exists (function (j,_,_) -> j=i) metasenv''')
            new_goals @ new_goals'
       in
       let res = proof'', patched_new_goals in
+      let upto = List.length actual_args in
       if upto = 0 then res else 
       let continuation = beta_after_elim_tac upto pred in
       let dummy_status = proof,goal in
@@ -714,9 +719,9 @@ let cases_intros_tac ?(howmany=(-1)) ?(mk_fresh_name_callback = FreshNamesGenera
   let module U = UriManager in
   let module R = CicReduction in
   let module C = Cic in
-  let (curi,metasenv,proofbo,proofty, attrs) = proof in
+  let (curi,metasenv,_subst, proofbo,proofty, attrs) = proof in
   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
-  let termty,_ = TC.type_of_aux' metasenv context term CicUniv.empty_ugraph in
+  let termty,_ = TC.type_of_aux' metasenv context term CicUniv.oblivion_ugraph in
   let termty = CicReduction.whd context termty in
   let (termty,metasenv',arguments,fresh_meta) =
    TermUtil.saturate_term
@@ -730,7 +735,7 @@ let cases_intros_tac ?(howmany=(-1)) ?(mk_fresh_name_callback = FreshNamesGenera
     | _ -> raise NotAnInductiveTypeToEliminate
   in
   let paramsno,itty,patterns,right_args =
-    match CicEnvironment.get_obj CicUniv.empty_ugraph uri with
+    match CicEnvironment.get_obj CicUniv.oblivion_ugraph uri with
     | C.InductiveDefinition (tys,_,paramsno,_),_ ->
        let _,left_parameters,right_args = 
          List.fold_right 
@@ -819,19 +824,19 @@ let cases_intros_tac ?(howmany=(-1)) ?(mk_fresh_name_callback = FreshNamesGenera
   let term_to_refine = C.MutCase (uri,typeno,outtype,term,patterns) in
   let refined_term,_,metasenv'',_ = 
     CicRefine.type_of_aux' metasenv' context term_to_refine
-      CicUniv.empty_ugraph
+      CicUniv.oblivion_ugraph
   in
   let new_goals =
     ProofEngineHelpers.compare_metasenvs
       ~oldmetasenv:metasenv ~newmetasenv:metasenv''
   in
-  let proof' = curi,metasenv'',proofbo,proofty, attrs in
+  let proof' = curi,metasenv'',_subst,proofbo,proofty, attrs in
   let proof'', new_goals' =
     PET.apply_tactic (apply_tac ~term:refined_term) (proof',goal)
   in
   (* The apply_tactic can have closed some of the new_goals *)
   let patched_new_goals =
-    let (_,metasenv''',_,_,_) = proof'' in
+    let (_,metasenv''',_subst,_,_,_) = proof'' in
       List.filter
         (function i -> List.exists (function (j,_,_) -> j=i) metasenv''')
         new_goals @ new_goals'
@@ -859,22 +864,3 @@ let elim_intros_simpl_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fres
        [ReductionTactics.simpl_tac
          ~pattern:(ProofEngineTypes.conclusion_pattern None)])
 ;;
-
-(* FG: insetrts a "hole" in the context (derived from letin_tac) *)
-
-let letout_tac =
-   let mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[] in
-   let term = C.Sort C.Set in
-   let letout_tac (proof, goal) =
-      let curi, metasenv, pbo, pty, attrs = proof in
-      let metano, context, ty = CicUtil.lookup_meta goal metasenv in
-      let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
-      let fresh_name = mk_fresh_name_callback metasenv context (Cic.Name "hole") ~typ:term in
-      let context_for_newmeta = None :: context in
-      let irl = CicMkImplicit.identity_relocation_list_for_metavariable context_for_newmeta in
-      let newmetaty = CicSubstitution.lift 1 ty in
-      let bo' = C.LetIn (fresh_name, term, C.Meta (newmeta,irl)) in
-      let newproof, _ = ProofEngineHelpers.subst_meta_in_proof proof metano bo'[newmeta,context_for_newmeta,newmetaty] in
-      newproof, [newmeta]
-   in
-   PET.mk_tactic letout_tac