(* Title: HOL/Library/conditional_parametricity.ML
Author: Jan Gilcher, Andreas Lochbihler, Dmitriy Traytel, ETH Zürich
A conditional parametricity prover
*)signatureCONDITIONAL_PARAMETRICITY=sigexception WARNING ofstringtypesettings={suppress_print_theorem:bool,suppress_warnings:bool,warnings_as_errors:bool,use_equality_heuristic:bool}val default_settings:settingsval quiet_settings:settingsval parametric_constant:settings->Attrib.binding * thm->Proof.context->(thm * Proof.context)val get_parametricity_theorems:Proof.context->thmlistval prove_goal:settings->Proof.context->thmoption->term->thmval prove_find_goal_cond:settings->Proof.context->thmlist->thmoption->term->thmval mk_goal:Proof.context->term->termval mk_cond_goal:Proof.context->thm->term * thmval mk_param_goal_from_eq_def:Proof.context->thm->termval step_tac:settings->Proof.context->thmlist->int->tacticendstructureConditional_Parametricity:CONDITIONAL_PARAMETRICITY=structtypesettings={suppress_print_theorem:bool,suppress_warnings:bool,warnings_as_errors:bool(* overrides suppress_warnings! *),use_equality_heuristic:bool};valquiet_settings={suppress_print_theorem=true,suppress_warnings=true,warnings_as_errors=false,use_equality_heuristic=false};valdefault_settings={suppress_print_theorem=false,suppress_warnings=false,warnings_as_errors=false,use_equality_heuristic=false};(* helper functions *)funstrip_imp_prems_concl(Const("Pure.imp",_)$A$B)=A::strip_imp_prems_conclB|strip_imp_prems_conclC=[C];funstrip_prop_safet=Logic.unprotectthandleTERM_=>t;funget_class_ofctxtt=Axclass.class_of_param(Proof_Context.theory_ofctxt)(fst(dest_Constt));funis_class_opctxtt=letvalt'=t|>Envir.eta_contract;inTerm.is_Constt'andalsois_some(get_class_ofctxtt')end;funapply_Var_to_boundst=letval(t,ts)=strip_combt;in(casetofVar(xi,_)=>letval(bounds,tail)=chop_prefixis_Boundts;inlist_comb(Var(xi,fastype_of(betapplys(t,bounds))),mapapply_Var_to_boundstail)end|_=>list_comb(t,mapapply_Var_to_boundsts))end;funtheorem_format_errorctxtthm=letvalmsg=Pretty.string_of(Pretty.chunks[(Pretty.para"Unexpected format of definition. Must be an unconditional equation."),Thm.pretty_thmctxtthm]);inerrormsgend;(* Tacticals and Tactics *)exceptionFINISHofthm;(* Tacticals *)funREPEAT_TRY_ELSE_DEFERtacst=letfunCOMB'taccountst=(letvaln=Thm.nprems_ofst;in(ifn=0thenall_tacstelse(caseSeq.pull((tacTHENCOMB'tac0)st)ofNONE=>ifcount+1=nthenraiseFINISHstelse(defer_tac1THEN(COMB'tac(count+1)))st|some=>Seq.make(fn()=>some)))end)inCOMB'tac0stend;(* Tactics *)(* helper tactics for printing *)funerror_tacctxtmsgst=(error(msg^"\n"^Goal_Display.string_of_goalctxtst);Seq.singlest);funerror_tac'ctxtmsg=SELECT_GOAL(error_tacctxtmsg);(* finds assumption of the form "Rel ?B Bound x Bound y", rotates it in front,
applies rel_app arity times and uses ams_rl *)funrel_app_tacctxttxyarity=letvalrel_app=[@{thmRel_app}];valassume=[asm_rl];funfind_and_rotate_tacti=letfunis_correct_rulet=(casetofConst(\<^const_name>‹HOL.Trueprop›,_)$(Const(\<^const_name>‹Transfer.Rel›,_)$_$Boundx'$Boundy')=>x=x'andalsoy=y'|_=>false);validx=find_indexis_correct_rule(t|>Logic.strip_assums_hyp);inifidx<0thenno_tacelserotate_tacidxiend;funrotate_and_dresolve_tacctxtarityi=REPEAT_DETERM_N(arity-1)(EVERY'[rotate_tac~1,dresolve_tacctxtrel_app,defer_tac]i);inSELECT_GOAL(EVERY'[find_and_rotate_tact,forward_tacctxtrel_app,defer_tac,rotate_and_dresolve_tacctxtarity,rotate_tac~1,eresolve_tacctxtassume]1)end;exceptionWARNINGofstring;funtransform_rules0thms=thms|transform_rulesnthms=transform_rules(n-1)(curry(Drule.RLoswap)@{thmsRel_appRel_match_app}thms);funassume_equality_tacsettingsctxttarityist=letvalquiet=#suppress_warningssettings;valerrors=#warnings_as_errorssettings;valT=fastype_oft;valis_eq_lemma=@{thmis_equality_Rel}|>Thm.incr_indexes((Term.maxidx_of_termt)+1)|>Drule.infer_instantiate'ctxt[NONE,SOME(Thm.cterm_ofctxtt)];valmsg=Pretty.string_of(Pretty.chunks[Pretty.paragraph((Pretty.text"No rule found for constant \"")@[Syntax.pretty_termctxtt,Pretty.str" :: ",Syntax.pretty_typctxtT]@(Pretty.text"\". Using is_eq_lemma:")),Pretty.quote(Thm.pretty_thmctxtis_eq_lemma)]);funmsg_tacst=(iferrorsthenraiseWARNINGmsgelseifquietthen()elsewarningmsg;Seq.singlest)valtac=resolve_tacctxt(transform_rulesarity[is_eq_lemma])i;in(iffold_atyps(K(Ktrue))Tfalsethenmsg_tacTHENtacelsetac)stend;funmark_class_as_match_tacctxtconstconst'arity=letvalrules=transform_rulesarity[@{thmRel_match_Rel}|>Thm.incr_indexes((Int.maxoapply2Term.maxidx_of_term)(const,const')+1)|>Drule.infer_instantiate'ctxt[NONE,SOME(Thm.cterm_ofctxtconst),SOME(Thm.cterm_ofctxtconst')]];inresolve_tacctxtrulesend;(* transforms the parametricity theorems to fit a given arity and uses them for resolution *)funparametricity_thm_tacsettingsctxtparametricity_thmsconstarity=letvalrules=transform_rulesarityparametricity_thms;inresolve_tacctxtrulesORELSE'assume_equality_tacsettingsctxtconstarityend;(* variant of parametricity_thm_tac to use matching *)funparametricity_thm_match_tacctxtparametricity_thmsarity=letvalrules=transform_rulesarityparametricity_thms;inmatch_tacctxtrulesend;funrel_abs_tacctxt=resolve_tacctxt[@{thmRel_abs}];funstep_tac'settingsctxtparametricity_thms(tm,i)=(casetm|>Logic.strip_assums_conclofConst(\<^const_name>‹HOL.Trueprop›,_)$(Const(rel,_)$_$t$u)=>letval(arity_of_t,arity_of_u)=apply2(strip_comb#>snd#>length)(t,u);in(caserelof\<^const_name>‹Transfer.Rel›=>(case(head_oft,head_ofu)of(Abs_,_)=>rel_abs_tacctxt|(_,Abs_)=>rel_abs_tacctxt|(constas(Const_),const'as(Const_))=>if#use_equality_heuristicsettingsandalsotaconvuthenassume_equality_tacquiet_settingsctxtt0elseifarity_of_t=arity_of_uthenifis_class_opctxtconstorelseis_class_opctxtconst'thenmark_class_as_match_tacctxtconstconst'arity_of_telseparametricity_thm_tacsettingsctxtparametricity_thmsconstarity_of_telseerror_tac'ctxt"Malformed term. Arities of t and u don't match."|(Boundx,Boundy)=>ifarity_of_t=arity_of_uthenifarity_of_t>0thenrel_app_tacctxttmxyarity_of_telseassume_tacctxtelseerror_tac'ctxt"Malformed term. Arities of t and u don't match."|_=>error_tac'ctxt"Unexpected format. Expected (Abs _, _), (_, Abs _), (Const _, Const _) or (Bound _, Bound _).")|\<^const_name>‹Conditional_Parametricity.Rel_match›=>parametricity_thm_match_tacctxtparametricity_thmsarity_of_t|_=>error_tac'ctxt"Unexpected format. Expected Transfer.Rel or Rel_match marker.")iend|Const(\<^const_name>‹HOL.Trueprop›,_)$(Const(\<^const_name>‹Transfer.is_equality›,_)$_)=>Transfer.eq_tacctxti|_=>error_tac'ctxt"Unexpected format. Not of form Const (HOL.Trueprop, _) $ _"i);funstep_tacsettings=SUBGOALoostep_tac'settings;funapply_theorem_tacctxtthm=HEADGOAL(resolve_tacctxt[Local_Defs.unfoldctxt@{thmsPure.prop_def}thm]THEN_ALL_NEWassume_tacctxt);(* Goal Generation *)funstrip_boundvars_from_rel_matcht=(casetof(TpasConst(\<^const_name>‹HOL.Trueprop›,_))$((RmasConst(\<^const_name>‹Conditional_Parametricity.Rel_match›,_))$R$t$t')=>Tp$(Rm$apply_Var_to_boundsR$t$t')|_=>t);valextract_conditions=letvalfilter_bounds=filter_outTerm.is_open;valprem_to_conditions=map(mapstrip_boundvars_from_rel_matchostrip_imp_prems_conclostrip_all_body);valremove_duplicates=distinctTerm.aconv;inremove_duplicatesofilter_boundsoflatoprem_to_conditionsend;funmk_goalctxtt=letvalctxt=fold(Variable.declare_typosnd)(Term.add_freest[])ctxt;valt=singleton(Variable.polymorphicctxt)t;vali=maxidx_of_termt+1;funtvar_to_tfree((name,_),sort)=(name,sort);valtvars=Term.add_tvarst[];valnew_frees=mapTFree(Term.variant_freest(maptvar_to_tfreetvars));valu=subst_atomic_types((mapTVartvars)~~new_frees)t;valT=fastype_oft;valU=fastype_ofu;valR=[T,U]--->\<^typ>‹bool›valr=Var(("R",2*i),R);valtransfer_rel=Const(\<^const_name>‹Transfer.Rel›,[R,T,U]--->\<^typ>‹bool›);inHOLogic.mk_Trueprop(transfer_rel$r$t$u)end;funmk_abs_helperTt=letvalU=fastype_oft;funmk_abs_helper'TU=ifT=Uthentelseletval(T2,T1)=Term.dest_funTT;inTerm.absdummyT2(mk_abs_helper'T1U)end;inmk_abs_helper'TUend;funcompare_ixs((name,i):indexname,(name',i'):indexname)=ifname<name'thenLESSelseifname>name'thenGREATERelseifi<i'thenLESSelseifi>i'thenGREATERelseEQUAL;funmk_cond_goalctxtthm=letvalconclusion=(hdostrip_imp_prems_conclostrip_prop_safeoThm.concl_of)thm;valconditions=(extract_conditionsoThm.prems_of)thm;valgoal=Logic.list_implies(conditions,conclusion);funcompare((ix,_),(ix',_))=compare_ixs(ix,ix');valgoal_vars=Term.add_varsgoal[]|>Ord_List.makecompare;val(ixs,Ts)=split_listgoal_vars;val(_,Ts')=Term.add_vars(Thm.prop_ofthm)[]|>Ord_List.makecompare|>Ord_List.intercomparegoal_vars|>split_list;val(As,_)=Ctr_Sugar_Util.mk_Frees"A"Tsctxt;valgoal_subst=ixs~~As;valthm_subst=ixs~~(map2mk_abs_helperTs'As);valthm'=thm|>Drule.infer_instantiatectxt(map(apsnd(Thm.cterm_ofctxt))thm_subst);in(goal|>Term.subst_Varsgoal_subst,thm')end;funmk_param_goal_from_eq_defctxtthm=letvalt=(caseThm.full_prop_ofthmof(Const(\<^const_name>‹Pure.eq›,_)$t'$_)=>t'|_=>theorem_format_errorctxtthm);inmk_goalctxttend;(* Transformations and parametricity theorems *)funtransform_class_rulectxtthm=(caseThm.concl_ofthmofConst(\<^const_name>‹HOL.Trueprop›,_)$(Const(\<^const_name>‹Transfer.Rel›,_)$_$t$u)=>(ifcurryTerm.aconv_untypedtuandalsois_class_opctxttthenthmRS@{thmRel_Rel_match}elsethm)|_=>thm);funis_parametricity_theoremthm=(caseThm.concl_ofthmofConst(\<^const_name>‹HOL.Trueprop›,_)$(Const(rel,_)$_$t$u)=>ifrel=\<^const_name>‹Transfer.Rel›orelserel=\<^const_name>‹Conditional_Parametricity.Rel_match›thencurryTerm.aconv_untypedtuelsefalse|_=>false);(* Pre- and postprocessing of theorems *)funmk_Domainp_assm(T,R)=HOLogic.mk_eq((Const(\<^const_name>‹Domainp›,Term.fastype_ofT-->Term.fastype_ofR)$T),R);valDomainp_lemma=@{lemma"(!!R.DomainpT=R==>PROP(PR))==PROP(P(DomainpT))"by(rule,drulemeta_spec,erulemeta_mp,ruleHOL.refl,simp)};funfold_Domainpf(tasConst(\<^const_name>‹Domainp›,_)$(Var(_,_)))=ft|fold_Domainpf(t$u)=fold_Domainpft#>fold_Domainpfu|fold_Domainpf(Abs(_,_,t))=fold_Domainpft|fold_Domainp__=I;funsubst_termstabt=letvalt'=Termtab.lookuptabtin(caset'ofSOMEt'=>t'|NONE=>(casetofu$v=>(subst_termstabu)$(subst_termstabv)|Abs(a,T,t)=>Abs(a,T,subst_termstabt)|t=>t))end;fungen_abstract_domainsctxt(dest:term->term * (term->term))thm=letvalprop=Thm.prop_ofthmval(t,mk_prop')=destpropvalDomainp_ts=rev(fold_Domainp(fnt=>insertop=t)t[])valDomainp_Ts=map(sndodest_funTosndodest_Constofstodest_comb)Domainp_tsvalused=Term.add_free_namest[]valrels=map(sndodest_comb)Domainp_tsvalrel_names=map(fstofstodest_Var)relsvalnames=map(fnname=>("D"^name))rel_names|>Name.variant_listusedvalfrees=mapFree(names~~Domainp_Ts)valprems=map(HOLogic.mk_Truepropomk_Domainp_assm)(rels~~frees);valt'=subst_terms(foldTermtab.update(Domainp_ts~~frees)Termtab.empty)tvalprop1=foldLogic.allfrees(Logic.list_implies(prems,mk_prop't'))valprop2=Logic.list_rename_params(revnames)prop1valcprop=Thm.cterm_ofctxtprop2valequal_thm=Raw_Simplifier.rewritectxtfalse[Domainp_lemma]cpropfunforall_elimthm=Thm.forall_elim_vars(Thm.maxidx_ofthm+1)thm;inforall_elim(thmCOMP(equal_thmCOMP@{thmequal_elim_rule2}))endhandleTERM_=>thm;funabstract_domains_transferctxtthm=letfundestprop=letvalprems=Logic.strip_imp_premspropvalconcl=HOLogic.dest_Trueprop(Logic.strip_imp_conclprop)val((rel,x),y)=apfstTerm.dest_comb(Term.dest_combconcl)in(x,fnx'=>Logic.list_implies(prems,HOLogic.mk_Trueprop(rel$x'$y)))endingen_abstract_domainsctxtdestthmend;funtransfer_rel_convconv=Conv.concl_conv~1(HOLogic.Trueprop_conv(Conv.fun2_conv(Conv.arg_convconv)));funfold_relator_eqs_convctxt=Conv.bottom_rewrs_conv(Transfer.get_relator_eqctxt)ctxt;funmk_is_equalityt=Const(\<^const_name>‹is_equality›,Term.fastype_oft-->HOLogic.boolT)$t;valis_equality_lemma=@{lemma"(!!R.is_equalityR==>PROP(PR))==PROP(P(=))"by(unfoldis_equality_def,rule,drulemeta_spec,erulemeta_mp,ruleHOL.refl,simp)};fungen_abstract_equalitiesctxt(dest:term->term * (term->term))thm=letvalprop=Thm.prop_ofthmval(t,mk_prop')=destprop(* Only consider "(=)" at non-base types *)funis_eq(Const(\<^const_name>‹HOL.eq›,Type("fun",[T,_])))=(caseTofType(_,[])=>false|_=>true)|is_eq_=falsevaladd_eqs=Term.fold_aterms(fnt=>ifis_eqttheninsert(op=)telseI)valeq_consts=rev(add_eqst[])valeqTs=map(sndodest_Const)eq_constsvalused=Term.add_free_namesprop[]valnames=map(K"")eqTs|>Name.variant_listusedvalfrees=mapFree(names~~eqTs)valprems=map(HOLogic.mk_Truepropomk_is_equality)freesvalprop1=mk_prop'(Term.subst_atomic(eq_consts~~frees)t)valprop2=foldLogic.allfrees(Logic.list_implies(prems,prop1))valcprop=Thm.cterm_ofctxtprop2valequal_thm=Raw_Simplifier.rewritectxtfalse[is_equality_lemma]cpropfunforall_elimthm=Thm.forall_elim_vars(Thm.maxidx_ofthm+1)thminforall_elim(thmCOMP(equal_thmCOMP@{thmequal_elim_rule2}))endhandleTERM_=>thm;funabstract_equalities_transferctxtthm=letfundestprop=letvalprems=Logic.strip_imp_premspropvalconcl=HOLogic.dest_Trueprop(Logic.strip_imp_conclprop)val((rel,x),y)=apfstTerm.dest_comb(Term.dest_combconcl)in(rel,fnrel'=>Logic.list_implies(prems,HOLogic.mk_Trueprop(rel'$x$y)))endvalcontracted_eq_thm=Conv.fconv_rule(transfer_rel_conv(fold_relator_eqs_convctxt))thmhandleCTERM_=>thmingen_abstract_equalitiesctxtdestcontracted_eq_thmend;funprep_rulectxt=abstract_equalities_transferctxt#>abstract_domains_transferctxt;funget_preprocess_theoremsctxt=Named_Theorems.getctxt\<^named_theorems>‹parametricity_preprocess›;funpreprocess_theoremctxt=Local_Defs.unfold0ctxt(get_preprocess_theoremsctxt)#>transform_class_rulectxt;funpostprocess_theoremctxt=Local_Defs.foldctxt(@{thmRel_Rel_match_eq}::get_preprocess_theoremsctxt)#>prep_rulectxt#>Local_Defs.unfoldctxt@{thmsRel_def};funget_parametricity_theoremsctxt=letvalparametricity_thm_map_filter=Option.filter(is_parametricity_theoremandf(notocurryTerm.could_unify(Thm.full_prop_of@{thmis_equality_Rel}))oThm.full_prop_of)opreprocess_theoremctxt;inmap_filter(parametricity_thm_map_filteroThm.transfer'ctxt)(Transfer.get_transfer_rawctxt)end;(* Provers *)(* Tries to prove a parametricity theorem without conditions, returns the last goal_state as thm *)funprove_find_goal_condsettingsctxtrulesdef_thmt=letfunfind_conditions_tac{context=ctxt,prems=_}=unfold_tacctxt(the_listdef_thm)THEN(REPEAT_TRY_ELSE_DEFERoHEADGOAL)(step_tacsettingsctxtrules);inGoal.provectxt[][]tfind_conditions_tachandleFINISHst=>stend;(* Simplifies and proves thm *)funprove_cond_goalctxtthm=letval(goal,thm')=mk_cond_goalctxtthm;valvars=Variable.add_free_namesctxtgoal[];funprove_conditions_tac{context=ctxt,prems=_}=apply_theorem_tacctxtthm';valvars=Variable.add_free_namesctxt(Thm.prop_ofthm')vars;inGoal.provectxtvars[]goalprove_conditions_tacend;(* Finds necessary conditions for t and proofs conditional parametricity of t under those conditions *)funprove_goalsettingsctxtdef_thmt=letvalparametricity_thms=get_parametricity_theoremsctxt;valfound_thm=prove_find_goal_condsettingsctxtparametricity_thmsdef_thmt;valthm=prove_cond_goalctxtfound_thm;inpostprocess_theoremctxtthmend;(* Commands *)fungen_parametric_constantsettingsprep_attprep_thm(raw_b:Attrib.binding,raw_eq)lthy=letvalb=apsnd(map(prep_attlthy))raw_b;valdef_thm=(prep_thmlthyraw_eq);valeq=Ctr_Sugar_Util.mk_abs_defdef_thmhandleTERM_=>theorem_format_errorlthydef_thm;valgoal=mk_param_goal_from_eq_deflthyeq;valthm=prove_goalsettingslthy(SOMEeq)goal;val(res,lthy')=Local_Theory.note(b,[thm])lthy;val_=if#suppress_print_theoremsettingsthen()elseProof_Display.print_resultstrue(Position.thread_data())lthy'(("theorem",""),[res]);in(the_single(sndres),lthy')end;funparametric_constantsettings=gen_parametric_constantsettings(KI)(KI);valparametric_constant_cmd=sndoogen_parametric_constantdefault_settings(Attrib.check_src)(singletonoAttrib.eval_thms);val_=Outer_Syntax.local_theory\<^command_keyword>‹parametric_constant›"proves parametricity"((Parse_Spec.opt_thm_name":"--Parse.thm)>>parametric_constant_cmd);end;