ProofEngineTypes.mk_tactic (discriminate'_tac ~term)
;;
+let exn_nonproj =
+ ProofEngineTypes.Fail (lazy "Injection: not a projectable equality");;
+let exn_noneq =
+ ProofEngineTypes.Fail (lazy "Injection: not an equality");;
+let exn_nothingtodo =
+ ProofEngineTypes.Fail (lazy "Nothing to do");;
+let exn_discrnonind =
+ ProofEngineTypes.Fail (lazy "Discriminate: object is not an Inductive Definition: it's imposible");;
+let exn_injwronggoal =
+ ProofEngineTypes.Fail (lazy "Injection: goal after cut is not correct");;
+let exn_noneqind =
+ ProofEngineTypes.Fail (lazy "Injection: not an equality over elements of an inductive type");;
+
let rec injection_tac ~first_time ~term ~liftno ~continuation =
+ let module C = Cic in
+ let module CR = CicReduction in
+ let module U = UriManager in
+ let module P = PrimitiveTactics in
+ let module T = Tacticals in
+ let module PST = ProofEngineStructuralRules in
+ let module PET = ProofEngineTypes in
+ let prerr_endline s = prerr_endline (String.make liftno ' ' ^ s) in
+ let are_convertible hd1 hd2 metasenv context =
+ fst (CR.are_convertible ~metasenv context hd1 hd2 CicUniv.empty_ugraph)
+ in
let injection_tac ~term status =
let (proof, goal) = status in
- let module C = Cic in
- let module U = UriManager in
- let module P = PrimitiveTactics in
- let module T = Tacticals in
let _,metasenv,_,_, _ = proof in
let _,context,_ = CicUtil.lookup_meta goal metasenv in
let term = CicSubstitution.lift liftno term in
- let termty,_ = (* TASSI: FIXME *)
+ let termty,_ =
CicTypeChecker.type_of_aux' metasenv context term CicUniv.empty_ugraph
in
- ProofEngineTypes.apply_tactic
- (match termty with
- (C.Appl [(C.MutInd (equri, 0, [])) ; tty ; t1 ; t2])
- when LibraryObjects.is_eq_URI equri ->
- begin
- match tty with
- (C.MutInd (turi,typeno,exp_named_subst))
- | (C.Appl (C.MutInd (turi,typeno,exp_named_subst)::_)) ->
- begin
- match t1,t2 with
- ((C.MutConstruct (uri1,typeno1,consno1,exp_named_subst1)),
- (C.MutConstruct (uri2,typeno2,consno2,exp_named_subst2)))
- when (uri1 = uri2) && (typeno1 = typeno2) &&
- (consno1 = consno2) &&
- (exp_named_subst1 = exp_named_subst2)
- ->
- if first_time then
- raise
- (ProofEngineTypes.Fail (lazy "Injection: nothing to do"))
- else
- continuation ~liftno
- | C.Appl
- ((C.MutConstruct
- (uri1,typeno1,consno1,exp_named_subst1))::applist1),
- C.Appl
- ((C.MutConstruct
- (uri2,typeno2,consno2,exp_named_subst2))::applist2)
- when (uri1 = uri2) && (typeno1 = typeno2) &&
- (consno1 = consno2) &&
- (exp_named_subst1 = exp_named_subst2)
- ->
- let rec traverse_list i l1 l2 =
- match l1,l2 with
- [],[] ->
- if first_time then
- continuation
- else
- (match term with
- C.Rel n ->
- (match List.nth context (n-1) with
- Some (C.Name id,_) ->
- fun ~liftno ->
- T.then_
- ~start:
- (ProofEngineStructuralRules.clear
- ~hyps:[id])
- ~continuation:(continuation ~liftno)
- | _ -> assert false)
- | _ -> assert false)
- | hd1::tl1,hd2::tl2 ->
- if
- fst
- (CicReduction.are_convertible ~metasenv
- context hd1 hd2 CicUniv.empty_ugraph)
- then
- traverse_list (i+1) tl1 tl2
- else
- injection1_tac ~i ~term
- ~continuation:(traverse_list (i+1) tl1 tl2)
- | _ -> assert false (* i 2 termini hanno in testa lo stesso costruttore, ma applicato a un numero diverso di termini *)
- in
- traverse_list 1 applist1 applist2 ~liftno
- | ((C.MutConstruct (uri1,typeno1,consno1,exp_named_subst1)),
- (C.MutConstruct (uri2,typeno2,consno2,exp_named_subst2)))
- | ((C.MutConstruct (uri1,typeno1,consno1,exp_named_subst1)),
- (C.Appl ((C.MutConstruct (uri2,typeno2,consno2,exp_named_subst2))::_)))
- | ((C.Appl ((C.MutConstruct (uri1,typeno1,consno1,exp_named_subst1))::_)),
- (C.MutConstruct (uri2,typeno2,consno2,exp_named_subst2)))
- | ((C.Appl ((C.MutConstruct (uri1,typeno1,consno1,exp_named_subst1))::_)),
- (C.Appl ((C.MutConstruct (uri2,typeno2,consno2,exp_named_subst2))::_)))
- when (consno1 <> consno2) || (exp_named_subst1 <> exp_named_subst2) ->
- discriminate_tac ~term
- (*raise (ProofEngineTypes.Fail (lazy "Injection: not a projectable equality but a discriminable one"))*)
- | _ ->
- if first_time then
- raise (ProofEngineTypes.Fail (lazy "Injection: not a projectable equality"))
- else
- continuation ~liftno
- end
- | _ ->
- if first_time then
- raise (ProofEngineTypes.Fail (lazy "Injection: not a projectable equality"))
- else
- continuation ~liftno
- end
- | _ -> raise (ProofEngineTypes.Fail (lazy "Injection: not an equation"))
- ) status
+ prerr_endline ("injection su: " ^ CicPp.ppterm termty);
+ let tac =
+ match termty with
+ | C.Appl [(C.MutInd (equri, 0, [])) ; tty ; t1 ; t2]
+ when LibraryObjects.is_eq_URI equri -> begin
+ match (CicReduction.whd ~delta:true context tty) with
+ | C.MutInd (turi,typeno,ens)
+ | C.Appl (C.MutInd (turi,typeno,ens)::_) -> begin
+ match t1,t2 with
+ | C.MutConstruct (uri1,typeno1,consno1,ens1),
+ C.MutConstruct (uri2,typeno2,consno2,ens2)
+ when (uri1 = uri2) && (typeno1 = typeno2) &&
+ (consno1 = consno2) && (ens1 = ens2) ->
+ if first_time then raise exn_nothingtodo
+ else continuation ~liftno
+ | C.Appl ((C.MutConstruct (uri1,typeno1,consno1,ens1))::applist1),
+ C.Appl ((C.MutConstruct (uri2,typeno2,consno2,ens2))::applist2)
+ when (uri1 = uri2) && (typeno1 = typeno2) &&
+ (consno1 = consno2) && (ens1 = ens2) ->
+ let rec traverse_list i l1 l2 =
+ match l1,l2 with
+ | [],[] when first_time -> continuation
+ | [],[] -> begin
+ match term with
+ | C.Rel n -> begin
+ match List.nth context (n-1) with
+ | Some (C.Name id,_) ->
+ fun ~liftno ->
+ T.then_ ~start:(PST.clear ~hyps:[id])
+ ~continuation:(continuation ~liftno)
+ | _ -> assert false
+ end
+ | _ -> assert false
+ end
+ | hd1::tl1,hd2::tl2 ->
+ if are_convertible hd1 hd2 metasenv context then
+ traverse_list (i+1) tl1 tl2
+ else
+ injection1_tac ~i ~term
+ ~continuation:(traverse_list (i+1) tl1 tl2)
+ | _ -> assert false
+ (* i 2 termini hanno in testa lo stesso costruttore,
+ * ma applicato a un numero diverso di termini *)
+ in
+ traverse_list 1 applist1 applist2 ~liftno
+ | C.MutConstruct (uri1,typeno1,consno1,ens1),
+ C.MutConstruct (uri2,typeno2,consno2,ens2)
+ | C.MutConstruct (uri1,typeno1,consno1,ens1),
+ C.Appl ((C.MutConstruct (uri2,typeno2,consno2,ens2))::_)
+ | C.Appl ((C.MutConstruct (uri1,typeno1,consno1,ens1))::_),
+ C.MutConstruct (uri2,typeno2,consno2,ens2)
+ | C.Appl ((C.MutConstruct (uri1,typeno1,consno1,ens1))::_),
+ C.Appl ((C.MutConstruct (uri2,typeno2,consno2,ens2))::_)
+ when (consno1 <> consno2) || (ens1 <> ens2) ->
+ discriminate_tac ~term
+ | _ when not first_time -> continuation ~liftno
+ | _ (* when first_time *) ->
+ match term with
+ | Cic.Rel i ->
+ let name =
+ match List.nth context (i-1) with
+ | Some (Cic.Name s, Cic.Def _) -> s
+ | Some (Cic.Name s, Cic.Decl _) -> s
+ | _ -> assert false
+ in
+ Tacticals.then_
+ ~start:(ReductionTactics.simpl_tac
+ ~pattern:(None,[name,Cic.Implicit (Some `Hole)],None))
+ ~continuation:(injection_tac ~first_time:false ~term ~liftno
+ ~continuation)
+ | _ -> raise exn_nonproj
+ end
+ | _ when not first_time -> continuation ~liftno
+ | _ (* when first_time *) -> raise exn_nonproj
+ end
+ | _ -> raise exn_nonproj
+ in
+ PET.apply_tactic tac status
in
- ProofEngineTypes.mk_tactic (injection_tac ~term)
+ PET.mk_tactic (injection_tac ~term)
and injection1_tac ~term ~i ~liftno ~continuation =
+ let module C = Cic in
+ let module CTC = CicTypeChecker in
+ let module CU = CicUniv in
+ let module S = CicSubstitution in
+ let module U = UriManager in
+ let module P = PrimitiveTactics in
+ let module PET = ProofEngineTypes in
+ let module T = Tacticals in
+ let prerr_endline s = prerr_endline (String.make liftno ' ' ^ s) in
+ let give_name seed = function
+ | C.Name _ as name -> name
+ | C.Anonymous -> C.Name (incr seed; "y" ^ string_of_int !seed)
+ in
+ let rec mk_rels = function | 0 -> [] | n -> C.Rel n :: (mk_rels (n - 1)) in
let injection1_tac ~term ~i status =
let (proof, goal) = status in
- (* precondizione: t1 e t2 hanno in testa lo stesso costruttore ma differiscono (o potrebbero differire?) nell'i-esimo parametro del costruttore *)
- let module C = Cic in
- let module S = CicSubstitution in
- let module U = UriManager in
- let module P = PrimitiveTactics in
- let module T = Tacticals in
- let term = CicSubstitution.lift liftno term in
- let _,metasenv,_,_, _ = proof in
- let _,context,_ = CicUtil.lookup_meta goal metasenv in
- let termty,_ = (* TASSI: FIXME *)
- CicTypeChecker.type_of_aux' metasenv context term CicUniv.empty_ugraph
- in
- match termty with (* an equality *)
- (C.Appl [(C.MutInd (equri, 0, [])) ; tty ; t1 ; t2])
- when LibraryObjects.is_eq_URI equri -> (
- match tty with (* some inductive type *)
- (C.MutInd (turi,typeno,exp_named_subst))
- | (C.Appl (C.MutInd (turi,typeno,exp_named_subst)::_)) ->
- let t1',t2',consno = (* sono i due sottotermini che differiscono *)
- match t1,t2 with
- ((C.Appl ((C.MutConstruct (uri1,typeno1,consno1,exp_named_subst1))::applist1)),
- (C.Appl ((C.MutConstruct (uri2,typeno2,consno2,exp_named_subst2))::applist2)))
- when (uri1 = uri2) && (typeno1 = typeno2) && (consno1 = consno2) && (exp_named_subst1 = exp_named_subst2) -> (* controllo ridondante *)
- (List.nth applist1 (i-1)),(List.nth applist2 (i-1)),consno2
- | _ -> assert false
- in
- let tty',_ =
- CicTypeChecker.type_of_aux' metasenv context t1'
- CicUniv.empty_ugraph in
- let patterns,outtype =
- match
- fst (CicEnvironment.get_obj CicUniv.empty_ugraph turi)
- with
- C.InductiveDefinition (ind_type_list,_,paramsno,_)->
- let _,_,_,constructor_list =
- List.nth ind_type_list typeno in
- let i_constr_id,_ =
- List.nth constructor_list (consno - 1) in
- let seed = ref 0 in
- let patterns =
- List.map
- (function (id,cty) ->
- let reduced_cty = CicReduction.whd context cty in
- let rec aux t k =
- match t with
- C.Prod (_,_,target) when k <= paramsno ->
- aux target (k+1)
- | C.Prod (binder,source,target) when k > paramsno ->
- let binder' =
- match binder with
- C.Name _ -> binder
- | C.Anonymous ->
- C.Name
- (incr seed; "y" ^ string_of_int !seed)
- in
- C.Lambda (binder',source,(aux target (k+1)))
- | _ ->
- let nr_param_constr = k - 1 - paramsno in
- if id = i_constr_id
- then C.Rel (k - i)
- else S.lift (nr_param_constr + 1) t1' (* + 1 per liftare anche il lambda aggiunto esternamente al case *)
- in aux reduced_cty 1
- ) constructor_list in
- let outtype =
- let seed = ref 0 in
- let rec to_lambdas te head =
- match CicReduction.whd context te with
- | C.Prod (binder,so,ta) ->
- let binder' =
- match binder with
- C.Name _ -> binder
- | C.Anonymous ->
- C.Name (incr seed; "d" ^ string_of_int !seed)
- in
- C.Lambda (binder',so,to_lambdas ta head)
- | _ -> head in
- let rec skip_prods n te =
- match n, CicReduction.whd context te with
- 0, _ -> te
- | n, C.Prod (_,_,ta) -> skip_prods (n - 1) ta
- | _, _ -> assert false
+ (* precondizione: t1 e t2 hanno in testa lo stesso costruttore ma
+ * differiscono (o potrebbero differire?) nell'i-esimo parametro
+ * del costruttore *)
+ let term = CicSubstitution.lift liftno term in
+ let _,metasenv,_,_, _ = proof in
+ let _,context,_ = CicUtil.lookup_meta goal metasenv in
+ let termty,_ =
+ CicTypeChecker.type_of_aux' metasenv context term CicUniv.empty_ugraph
+ in
+ prerr_endline ("injection1 su : " ^ CicPp.ppterm termty);
+ match termty with (* an equality *)
+ | C.Appl [(C.MutInd (equri, 0, [])) ; tty ; t1 ; t2]
+ when LibraryObjects.is_eq_URI equri ->
+ let turi,typeno,ens,params =
+ match tty with (* some inductive type *)
+ | C.MutInd (turi,typeno,ens) -> turi,typeno,ens,[]
+ | C.Appl (C.MutInd (turi,typeno,ens)::params) -> turi,typeno,ens,params
+ | _ -> raise exn_noneqind
+ in
+ let t1',t2',consno = (* sono i due sottotermini che differiscono *)
+ match t1,t2 with
+ | C.Appl ((C.MutConstruct (uri1,typeno1,consno1,ens1))::applist1),
+ C.Appl ((C.MutConstruct (uri2,typeno2,consno2,ens2))::applist2)
+ when (uri1 = uri2) && (typeno1 = typeno2) &&
+ (consno1 = consno2) && (ens1 = ens2) ->
+ (* controllo ridondante *)
+ List.nth applist1 (i-1),List.nth applist2 (i-1),consno2
+ | _ -> assert false
+ in
+ let tty',_ = CTC.type_of_aux' metasenv context t1' CU.empty_ugraph in
+ let patterns,outtype =
+ match fst (CicEnvironment.get_obj CicUniv.empty_ugraph turi) with
+ | C.InductiveDefinition (ind_type_list,_,paramsno,_)->
+ let left_params, right_params = HExtlib.split_nth paramsno params in
+ let _,_,_,constructor_list = List.nth ind_type_list typeno in
+ let i_constr_id,_ = List.nth constructor_list (consno - 1) in
+ let patterns =
+ let seed = ref 0 in
+ List.map
+ (function (id,cty) ->
+ let reduced_cty = CicReduction.whd context cty in
+ let rec aux k = function
+ | C.Prod (_,_,tgt) when k <= paramsno ->
+ let left = List.nth left_params (k-1) in
+ aux (k+1) (CicSubstitution.subst left tgt)
+ | C.Prod (binder,source,target) when k > paramsno ->
+ let binder' = give_name seed binder in
+ C.Lambda (binder',source,(aux (k+1) target))
+ | _ ->
+ let nr_param_constr = k - paramsno in
+ if id = i_constr_id then C.Rel (k - i)
+ else S.lift nr_param_constr t1'
+ (* + 1 per liftare anche il lambda aggiunto
+ * esternamente al case *)
+ in aux 1 reduced_cty)
+ constructor_list
+ in
+ let outtype =
+ let seed = ref 0 in
+ let rec to_lambdas te head =
+ match CicReduction.whd context te with
+ | C.Prod (binder,so,ta) ->
+ let binder' = give_name seed binder in
+ C.Lambda (binder',so,to_lambdas ta head)
+ | _ -> head
+ in
+ let rec skip_prods params te =
+ match params, CicReduction.whd context te with
+ | [], _ -> te
+ | left::tl, C.Prod (_,_,ta) ->
+ skip_prods tl (CicSubstitution.subst left ta)
+ | _, _ -> assert false
+ in
+ let abstracted_tty =
+ match CicSubstitution.lift (paramsno(* + 1*)) tty with
+ | C.MutInd _ as tty' -> tty'
+ | C.Appl l ->
+ let keep,abstract = HExtlib.split_nth (paramsno +1) l in
+ C.Appl (keep@mk_rels (List.length abstract))
+ | _ -> assert false
+ in
+ match ind_type_list with
+ | [] -> assert false
+ | (_,_,ty,_)::_ ->
+ to_lambdas (skip_prods left_params ty)
+ (C.Lambda
+ (C.Name "x", abstracted_tty, S.lift 1 tty'))
+ in
+ patterns,outtype
+ | _ -> raise exn_discrnonind
+ in
+ let cutted = C.Appl [C.MutInd (equri,0,[]) ; tty' ; t1' ; t2'] in
+ prerr_endline ("CUT: " ^ CicPp.ppterm cutted);
+ PET.apply_tactic
+ (T.thens ~start: (P.cut_tac cutted)
+ ~continuations:
+ [injection_tac ~first_time:false ~liftno:0 ~term:(C.Rel 1)
+ ~continuation:
+ (fun ~liftno:x -> continuation ~liftno:(liftno+1+x))
+ (* here I need to lift all the continuations by 1;
+ since I am setting back liftno to 0, I actually
+ need to lift all the continuations by liftno + 1 *)
+ ;T.then_
+ ~start:(PET.mk_tactic
+ (fun status ->
+ prerr_endline "riempo il cut";
+ let (proof, goal) = status in
+ let _,metasenv,_,_, _ = proof in
+ let _,context,gty =CicUtil.lookup_meta goal metasenv in
+ let gty = Unshare.unshare gty in
+ let new_t1' =
+ match gty with
+ | (C.Appl (C.MutInd (_,_,_)::_::t::_)) -> t
+ | _ -> raise exn_injwronggoal
in
- let abstracted_tty =
- match CicSubstitution.lift (paramsno + 1) tty with
- C.MutInd _ as tty' -> tty'
- | C.Appl l ->
- let keep,abstract =
- HExtlib.split_nth (paramsno +1) l in
- let rec mk_rels =
- function
- 0 -> []
- | n -> C.Rel n :: (mk_rels (n - 1))
- in
- C.Appl (keep@mk_rels (List.length abstract))
- | _ -> assert false
- in
- match ind_type_list with
- [] -> assert false
- | (_,_,ty,_)::_ ->
- to_lambdas (skip_prods paramsno ty)
- (C.Lambda (C.Name "x", abstracted_tty,
- S.lift (2+paramsno) tty'))
+ let changed =
+ C.Appl [
+ C.Lambda (C.Name "x", tty,
+ C.MutCase (turi,typeno,outtype,C.Rel 1,patterns))
+ ; t1]
in
- patterns,outtype
- | _ -> raise (ProofEngineTypes.Fail (lazy "Discriminate: object is not an Inductive Definition: it's imposible"))
- in
- ProofEngineTypes.apply_tactic
- (T.thens
- ~start:
- (P.cut_tac
- (C.Appl [C.MutInd (equri,0,[]) ; tty' ; t1' ; t2']))
- ~continuations:
- [ injection_tac ~first_time:false ~liftno:0
- ~term:(C.Rel 1)
- (* here I need to lift all the continuations by 1;
- since I am setting back liftno to 0, I actually
- need to lift all the continuations by liftno + 1 *)
- ~continuation:
- (fun ~liftno:x ->
- continuation ~liftno:(liftno + 1 + x)) ;
- T.then_
- ~start:(ProofEngineTypes.mk_tactic
- (fun status ->
- let (proof, goal) = status in
- let _,metasenv,_,_, _ = proof in
- let _,context,gty =
- CicUtil.lookup_meta goal metasenv
- in
- let new_t1' =
- match gty with
- (C.Appl (C.MutInd (_,_,_)::arglist)) ->
- List.nth arglist 1
- | _ ->
- raise
- (ProofEngineTypes.Fail
- (lazy
- "Injection: goal after cut is not correct"))
- in
- ProofEngineTypes.apply_tactic
- (ReductionTactics.change_tac
- ~pattern:(ProofEngineTypes.conclusion_pattern
- (Some new_t1'))
- (fun _ m u ->
- C.Appl [
- C.Lambda
- (C.Name "x",
- tty,
- C.MutCase
- (turi,typeno,outtype,C.Rel 1,patterns)) ;
- t1],
- m, u))
+ prerr_endline
+ ("metto questo: " ^ CicPp.ppterm changed);
+ prerr_endline
+ ("al posto di questo: " ^ CicPp.ppterm new_t1');
+ prerr_endline
+ ("nel goal: " ^ CicPp.ppterm gty);
+ prerr_endline
+ ("nel contesto:\n" ^ CicPp.ppcontext context);
+ let rc =
+ PET.apply_tactic
+ (ReductionTactics.change_tac
+ ~pattern:(None, [],
+ Some (ProofEngineHelpers.pattern_of
+ ~term:gty [new_t1']))
+ (fun _ m u -> changed,m,u))
status
- ))
- ~continuation:
- (T.then_
- ~start:
- (EqualityTactics.rewrite_simpl_tac
- ~direction:`LeftToRight
- ~pattern:(ProofEngineTypes.conclusion_pattern None)
- term [])
- ~continuation:EqualityTactics.reflexivity_tac)
- ])
- status
- | _ -> raise (ProofEngineTypes.Fail (lazy "Injection: not an equality over elements of an inductive type"))
- )
- | _ -> raise (ProofEngineTypes.Fail (lazy "Injection: not an equality"))
+ in prerr_endline "fine";rc
+ ))
+ ~continuation:
+ (T.then_
+ ~start:
+ (EqualityTactics.rewrite_simpl_tac
+ ~direction:`LeftToRight
+ ~pattern:(PET.conclusion_pattern None)
+ term [])
+ ~continuation:EqualityTactics.reflexivity_tac)
+ ])
+ status
+ | _ -> raise exn_noneq
in
- ProofEngineTypes.mk_tactic (injection1_tac ~term ~i)
+ PET.mk_tactic (injection1_tac ~term ~i)
;;
(* destruct performs either injection or discriminate *)