LamEx.thy
changeset 253 e169a99c6ada
parent 252 e30997c88050
child 257 68bd5c2a1b96
equal deleted inserted replaced
252:e30997c88050 253:e169a99c6ada
   194 thm alpha.cases
   194 thm alpha.cases
   195 thm alpha_cases
   195 thm alpha_cases
   196 
   196 
   197 thm rlam.inject
   197 thm rlam.inject
   198 thm pi_var
   198 thm pi_var
       
   199  
   199 
   200 
   200 lemma var_inject:
   201 lemma var_inject:
   201   shows "(Var a = Var b) = (a = b)"
   202   shows "(Var a = Var b) = (a = b)"
   202 sorry
   203 sorry
   203 
   204 
   219 
   220 
   220 
   221 
   221 
   222 
   222 
   223 
   223 
   224 
   224 ML {* val t_a = atomize_thm @{thm alpha.cases} *}
   225 ML {* val t_a = atomize_thm @{thm alpha.induct} *}
       
   226 (* prove {* build_regularize_goal t_a rty rel @{context}  *}
       
   227 ML_prf {*  fun tac ctxt =
       
   228      (FIRST' [
       
   229       rtac rel_refl,
       
   230       atac,
       
   231       rtac @{thm universal_twice},
       
   232       (rtac @{thm impI} THEN' atac),
       
   233       rtac @{thm implication_twice},
       
   234       EqSubst.eqsubst_tac ctxt [0]
       
   235         [(@{thm equiv_res_forall} OF [rel_eqv]),
       
   236          (@{thm equiv_res_exists} OF [rel_eqv])],
       
   237       (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
       
   238       (rtac @{thm RIGHT_RES_FORALL_REGULAR})
       
   239      ]);
       
   240  *}
       
   241   apply (tactic {* tac @{context} 1 *}) *)
   225 ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
   242 ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
   226 ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
   243 ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
   227 ML {* val abs = findabs rty (prop_of t_a); *}
   244 
   228 ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *}
   245 ML {* val abs = findabs rty (prop_of (atomize_thm @{thm alpha.induct})) *}
   229 ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
   246 ML {* val aps = findaps rty (prop_of (atomize_thm @{thm alpha.induct})) *}
       
   247 ML {* val aps = @{typ "LamEx.rlam \<Rightarrow> bool"} :: aps; *}
       
   248 ML {* val simp_app_prs_thms = map (make_simp_prs_thm @{context} quot @{thm APP_PRS}) aps *}
       
   249 ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
       
   250 ML {* val t_l = repeat_eqsubst_thm @{context} (simp_app_prs_thms @ simp_lam_prs_thms) t_t *}
   230 ML {* val t_a = simp_allex_prs @{context} quot t_l *}
   251 ML {* val t_a = simp_allex_prs @{context} quot t_l *}
       
   252 ML {* val thm = 
       
   253   @{thm FORALL_PRS[OF FUN_QUOTIENT[OF QUOTIENT_lam FUN_QUOTIENT[OF QUOTIENT_lam IDENTITY_QUOTIENT]], symmetric]} *}
       
   254 ML {* val t_a1 = eqsubst_thm @{context} [thm] t_a *}
   231 ML {* val defs_sym = add_lower_defs @{context} defs; *}
   255 ML {* val defs_sym = add_lower_defs @{context} defs; *}
   232 ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_a *}
   256 ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_a1 *}
   233 ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
   257 ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
   234 ML {* ObjectLogic.rulify t_r *}
   258 ML {* val t_r1 = repeat_eqsubst_thm @{context} @{thms fun_map.simps} t_r *}
   235 
   259 ML {* val t_r2 = repeat_eqsubst_thm @{context} @{thms QUOT_TYPE_I_lam.thm10} t_r1 *}
   236 
   260 ML {* val t_r3 = repeat_eqsubst_thm @{context} @{thms id_apply} t_r2 *}
       
   261 
       
   262 ML {* val alpha_induct = ObjectLogic.rulify t_r3 *}
       
   263 
       
   264 local_setup {*
       
   265   Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd
       
   266 *}
       
   267 
       
   268 lemma
       
   269   assumes a: "(a::lam) = b"
       
   270   shows "False"
       
   271   using a
       
   272   apply (rule alpha_induct)
   237 
   273 
   238 fun
   274 fun
   239   option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
   275   option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
   240 where
   276 where
   241   "option_map f (nSome x) = nSome (f x)"
   277   "option_map f (nSome x) = nSome (f x)"