~naesten/coq/trunk

« back to all changes in this revision

Viewing changes to plugins/funind/indfun.ml

  • Committer: herbelin
  • Date: 2009-11-08 17:31:16 UTC
  • Revision ID: svn-v4:85f007b7-540e-0410-9357-904b9bb8a0f7:trunk:12481
Restructuration of command.ml + generic infrastructure for inductive schemes
- Cleaning and uniformisation in command.ml:
  - For better modularity and better visibility, two files got isolated
    out of command.ml:
    - lemmas.ml is about starting and saving a proof
    - indschemes.ml is about declaring inductive schemes
  - Decomposition of the functions of command.ml into a functional part
    and the imperative part
- Inductive schemes:
  - New architecture in ind_tables.ml for registering scheme builders,
    and for sharing and generating on demand inductive schemes
  - Adding new automatically generated equality schemes (file eqschemes.ml)
    - "_congr" for equality types (completing here commit 12273)
    - "_rew_forward" (similar to vernac-level eq_rect_r), "_rew_forward_dep",
      "_rew_backward" (similar to eq_rect), "_rew_backward_dep" for
      rewriting schemes (warning, rew_forward_dep cannot be stated following
      the standard Coq pattern for inductive types: "t=u" cannot be the
      last argument of the scheme)
    - "_case", "_case_nodep", "_case_dep" for case analysis schemes
  - Preliminary step towards discriminate and injection working on any
    equality-like type (e.g. eq_true)
  - Restating JMeq_congr under the canonical form of congruence schemes
  - Renamed "Set Equality Scheme" into "Set Equality Schemes"
  - Added "Set Rewriting Schemes", "Set Case Analysis Schemes"
  - Activation of the automatic generation of boolean equality lemmas
  - Partial debug and error messages improvements for the generation of
    boolean equality and decidable equality
  - Added schemes for making dependent rewrite working (unfortunately with
    not a fully satisfactory design - see file eqschemes.ml)
- Some names of ML function made more regular (see dev/doc/changes.txt)
- Incidentally, added a flush to obsolete Local/Global syntax warning

Show diffs side-by-side

added added

removed removed

Lines of Context:
150
150
 
151
151
let interp_casted_constr_with_implicits sigma env impls c  =
152
152
(*   Constrintern.interp_rawconstr_with_implicits sigma env [] impls c *)
153
 
  Constrintern.intern_gen false sigma env ~impls:([],impls)
 
153
  Constrintern.intern_gen false sigma env ~impls
154
154
    ~allow_patvar:false  ~ltacvars:([],[]) c
155
155
 
156
156
 
166
166
  let (rec_sign,rec_impls) =
167
167
    List.fold_left
168
168
      (fun (env,impls) ((_,recname),_,bl,arityc,_) ->
169
 
        let arityc = Command.generalize_constr_expr arityc bl in
 
169
        let arityc = Topconstr.prod_constr_expr arityc bl in
170
170
        let arity = Constrintern.interp_type sigma env0 arityc in
171
 
        let impl =
172
 
          if Impargs.is_implicit_args()
173
 
          then Impargs.compute_implicits  env0 arity
174
 
          else [] in
175
 
        let impls' =(recname,(Constrintern.Recursive,[],impl,Notation.compute_arguments_scope arity))::impls in
176
 
        (Environ.push_named (recname,None,arity) env, impls'))
 
171
        let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity [] in
 
172
        (Environ.push_named (recname,None,arity) env, (recname,impl) :: impls))
177
173
      (env0,[]) lnameargsardef in
 
174
  let rec_impls = Constrintern.set_internalization_env_params rec_impls [] in
178
175
  let recdef =
179
176
    (* Declare local notations *)
180
177
    let fs = States.freeze() in
367
364
let register_struct is_rec fixpoint_exprl =
368
365
  match fixpoint_exprl with
369
366
    | [((_,fname),_,bl,ret_type,body),_] when not is_rec ->
 
367
        let ce,imps =
 
368
          Command.interp_definition
 
369
            (Flags.boxed_definitions ()) bl None body (Some ret_type)
 
370
        in
370
371
        Command.declare_definition
371
 
          fname
372
 
          (Decl_kinds.Global,Flags.boxed_definitions (),Decl_kinds.Definition)
373
 
          bl
374
 
          None
375
 
          body
376
 
          (Some ret_type)
377
 
          (fun _ _ -> ())
 
372
          fname (Decl_kinds.Global,Decl_kinds.Definition)
 
373
          ce imps (fun _ _ -> ())
378
374
    | _ ->
379
 
        Command.build_recursive fixpoint_exprl (Flags.boxed_definitions())
 
375
        Command.do_fixpoint fixpoint_exprl (Flags.boxed_definitions())
380
376
 
381
377
let generate_correction_proof_wf f_ref tcc_lemma_ref
382
378
    is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation
389
385
let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas args ret_type body
390
386
    pre_hook
391
387
    =
392
 
  let type_of_f = Command.generalize_constr_expr ret_type args in
 
388
  let type_of_f = Topconstr.prod_constr_expr ret_type args in
393
389
  let rec_arg_num =
394
390
    let names =
395
391
      List.map
420
416
    Topconstr.CApp (dummy_loc,(None,Topconstr.mkRefC (Qualid (dummy_loc,(qualid_of_string "Logic.eq")))),
421
417
                    [(f_app_args,None);(body,None)])
422
418
  in
423
 
  let eq = Command.generalize_constr_expr unbounded_eq args in
 
419
  let eq = Topconstr.prod_constr_expr unbounded_eq args in
424
420
  let hook f_ref tcc_lemma_ref functional_ref eq_ref rec_arg_num rec_arg_type
425
421
      nb_args relation =
426
422
    try
531
527
                         raise (UserError("",str "Cannot find argument " ++
532
528
                                            Ppconstr.pr_id id))
533
529
                     in
534
 
                     (name,annot,args,types,body),(None:Vernacexpr.decl_notation)
 
530
                     (name,annot,args,types,body),(None:Vernacexpr.decl_notation option)
535
531
                 | (name,None,args,types,body),recdef ->
536
532
                     let names =  (Topconstr.names_of_local_assums args) in
537
533
                     if  is_one_rec recdef  && List.length names > 1 then
541
537
                     else
542
538
                       let loc, na = List.hd names in
543
539
                         (name,(Some (loc, Nameops.out_name na), Topconstr.CStructRec),args,types,body),
544
 
                     (None:Vernacexpr.decl_notation)
 
540
                     (None:Vernacexpr.decl_notation option)
545
541
                 | (_,Some (Wf _),_,_,_),_ | (_,Some (Mes _),_,_,_),_->
546
542
                     error
547
543
                       ("Cannot use mutual definition with well-founded recursion or measure")