LamEx.thy
changeset 269 fe6eb116b341
parent 267 3764566c1151
child 270 c55883442514
equal deleted inserted replaced
267:3764566c1151 269:fe6eb116b341
   164 
   164 
   165 ML {* val qty = @{typ "lam"} *}
   165 ML {* val qty = @{typ "lam"} *}
   166 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
   166 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
   167 ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
   167 ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
   168 ML {* val consts = lookup_quot_consts defs *}
   168 ML {* val consts = lookup_quot_consts defs *}
   169 ML {* val (trans2, reps_same, quot) = lookup_quot_thms @{context} "lam" *}
   169 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
   170 ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
   170 ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
   171 ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
   171 ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
   172 
   172 
   173 ML {* val pi_var = lift_thm_lam @{context} @{thm pi_var_com} *}
   173 ML {* val pi_var = lift_thm_lam @{context} @{thm pi_var_com} *}
   174 ML {* val pi_app = lift_thm_lam @{context} @{thm pi_app_com} *}
   174 ML {* val pi_app = lift_thm_lam @{context} @{thm pi_app_com} *}