Playing with alpha_refl.
theory LamEx
imports Nominal QuotMain
begin
atom_decl name
thm abs_fresh(1)
nominal_datatype rlam =
rVar "name"
| rApp "rlam" "rlam"
| rLam "name" "rlam"
print_theorems
function
rfv :: "rlam \<Rightarrow> name set"
where
rfv_var: "rfv (rVar a) = {a}"
| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
sorry
termination rfv sorry
inductive
alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
where
a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
| a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
print_theorems
lemma alpha_refl:
shows "x \<approx> x"
apply (rule rlam.induct)
apply (simp_all add:a1 a2)
apply (rule a3)
apply (simp_all)
sorry
quotient lam = rlam / alpha
sorry
print_quotients
quotient_def
Var :: "name \<Rightarrow> lam"
where
"Var \<equiv> rVar"
quotient_def
App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
where
"App \<equiv> rApp"
quotient_def
Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
where
"Lam \<equiv> rLam"
thm Var_def
thm App_def
thm Lam_def
quotient_def
fv :: "lam \<Rightarrow> name set"
where
"fv \<equiv> rfv"
thm fv_def
(* definition of overloaded permutation function *)
(* for the lifted type lam *)
overloading
perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam" (unchecked)
begin
quotient_def
perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
where
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
end
(*quotient_def (for lam)
abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
where
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*)
thm perm_lam_def
(* lemmas that need to lift *)
lemma pi_var_com:
fixes pi::"'x prm"
shows "(pi\<bullet>rVar a) \<approx> rVar (pi\<bullet>a)"
sorry
lemma pi_app_com:
fixes pi::"'x prm"
shows "(pi\<bullet>rApp t1 t2) \<approx> rApp (pi\<bullet>t1) (pi\<bullet>t2)"
sorry
lemma pi_lam_com:
fixes pi::"'x prm"
shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)"
sorry
lemma fv_var:
shows "fv (Var a) = {a}"
sorry
lemma fv_app:
shows "fv (App t1 t2) = (fv t1) \<union> (fv t2)"
sorry
lemma fv_lam:
shows "fv (Lam a t) = (fv t) - {a}"
sorry
lemma real_alpha:
assumes "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
shows "Lam a t = Lam b s"
sorry
lemma perm_rsp: "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
apply(auto)
(* this is propably true if some type conditions are imposed ;o) *)
sorry
lemma fresh_rsp: "(op = ===> alpha ===> op =) fresh fresh"
apply(auto)
(* this is probably only true if some type conditions are imposed *)
sorry
lemma rVar_rsp: "(op = ===> alpha) rVar rVar"
apply(auto)
apply(rule a1)
apply(simp)
done
lemma rApp_rsp: "(alpha ===> alpha ===> alpha) rApp rApp"
apply(auto)
apply(rule a2)
apply (assumption)
apply (assumption)
done
lemma rLam_rsp: "(op = ===> alpha ===> alpha) rLam rLam"
apply(auto)
apply(rule a3)
apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst)
apply(rule sym)
apply(rule trans)
apply(rule pt_name3)
apply(rule at_ds1[OF at_name_inst])
apply(simp add: pt_name1)
apply(assumption)
apply(simp add: abs_fresh)
done
lemma rfv_rsp: "(alpha ===> op =) rfv rfv"
sorry
ML {* val qty = @{typ "lam"} *}
ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} @
@{thms ho_all_prs ho_ex_prs} *}
ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
ML {* val pi_var = lift_thm_lam @{context} @{thm pi_var_com} *}
ML {* val pi_app = lift_thm_lam @{context} @{thm pi_app_com} *}
ML {* val pi_lam = lift_thm_lam @{context} @{thm pi_lam_com} *}
ML {* val fv_var = lift_thm_lam @{context} @{thm rfv_var} *}
ML {* val fv_app = lift_thm_lam @{context} @{thm rfv_app} *}
ML {* val fv_lam = lift_thm_lam @{context} @{thm rfv_lam} *}
ML {* val a1 = lift_thm_lam @{context} @{thm a1} *}
ML {* val a2 = lift_thm_lam @{context} @{thm a2} *}
ML {* val a3 = lift_thm_lam @{context} @{thm a3} *}
ML {* val alpha_cases = lift_thm_lam @{context} @{thm alpha.cases} *}
ML {* val alpha_induct = lift_thm_lam @{context} @{thm alpha.induct} *}
local_setup {*
Quotient.note (@{binding "pi_var"}, pi_var) #> snd #>
Quotient.note (@{binding "pi_app"}, pi_app) #> snd #>
Quotient.note (@{binding "pi_lam"}, pi_lam) #> snd #>
Quotient.note (@{binding "a1"}, a1) #> snd #>
Quotient.note (@{binding "a2"}, a2) #> snd #>
Quotient.note (@{binding "a3"}, a3) #> snd #>
Quotient.note (@{binding "alpha_cases"}, alpha_cases) #> snd #>
Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd
*}
thm alpha.cases
thm alpha_cases
thm alpha.induct
thm alpha_induct
lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)"
apply (auto)
apply (erule alpha.cases)
apply (simp_all add: rlam.inject alpha_refl)
done
ML {* val var_inject = Toplevel.program (fn () => lift_thm_lam @{context} @{thm rvar_inject}) *}
local_setup {*
Quotient.note (@{binding "var_inject"}, var_inject) #> snd
*}
lemma var_supp:
shows "supp (Var a) = ((supp a)::name set)"
apply(simp add: supp_def)
apply(simp add: pi_var)
apply(simp add: var_inject)
done
lemma var_fresh:
fixes a::"name"
shows "(a\<sharp>(Var b)) = (a\<sharp>b)"
apply(simp add: fresh_def)
apply(simp add: var_supp)
done
(* Construction Site code *)
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
ML {* val consts = lookup_quot_consts defs *}
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
ML {* val t_a = atomize_thm @{thm alpha.induct} *}
(* prove {* build_regularize_goal t_a rty rel @{context} *}
ML_prf {* fun tac ctxt =
(FIRST' [
rtac rel_refl,
atac,
rtac @{thm universal_twice},
(rtac @{thm impI} THEN' atac),
rtac @{thm implication_twice},
EqSubst.eqsubst_tac ctxt [0]
[(@{thm equiv_res_forall} OF [rel_eqv]),
(@{thm equiv_res_exists} OF [rel_eqv])],
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
]);
*}
apply (tactic {* tac @{context} 1 *}) *)
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
(*ML {*
val rt = build_repabs_term @{context} t_r consts rty qty
val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
*}
prove rg
apply(atomize(full))
ML_prf {*
fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms =
(FIRST' [
rtac trans_thm,
LAMBDA_RES_TAC ctxt,
res_forall_rsp_tac ctxt,
res_exists_rsp_tac ctxt,
(
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps rsp_thms))
THEN_ALL_NEW (fn _ => no_tac)
),
(instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
rtac refl,
(APPLY_RSP_TAC rty ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])),
Cong_Tac.cong_tac @{thm cong},
rtac @{thm ext},
rtac reflex_thm,
atac,
(
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
THEN_ALL_NEW (fn _ => no_tac)
),
WEAK_LAMBDA_RES_TAC ctxt
]);
fun r_mk_comb_tac_lam lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms
*}
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_lam @{context}) 1 *})
*)
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
ML {* val abs = findabs rty (prop_of (atomize_thm @{thm alpha.induct})) *}
ML {* val aps = findaps rty (prop_of (atomize_thm @{thm alpha.induct})) *}
ML {* val (alls, exs) = findallex rty qty (prop_of (atomize_thm @{thm alpha.induct})) *}
ML {* val allthms = map (make_allex_prs_thm @{context} quot @{thm FORALL_PRS} ) alls *}
ML {* val exthms = map (make_allex_prs_thm @{context} quot @{thm EXISTS_PRS} ) exs *}
ML {* val t_a = MetaSimplifier.rewrite_rule allthms t_t *}
ML {* val simp_app_prs_thms = map (make_simp_prs_thm @{context} quot @{thm APP_PRS}) aps *}
ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
ML {* val t_l = repeat_eqsubst_thm @{context} (simp_lam_prs_thms) t_a *}
ML {* val t_l1 = repeat_eqsubst_thm @{context} simp_app_prs_thms t_l *}
ML {* val defs_sym = add_lower_defs @{context} defs; *}
ML {* val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym *}
ML {* val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_l1 *}
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_d0 *}
ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
ML {* val t_r1 = repeat_eqsubst_thm @{context} @{thms fun_map.simps} t_r *}
ML {* val t_r2 = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_lam.thm10} t_r1 *}
ML {* val t_r3 = MetaSimplifier.rewrite_rule @{thms eq_reflection[OF id_apply]} t_r2 *}
ML {* val alpha_induct = ObjectLogic.rulify t_r3 *}
(*local_setup {*
Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd
*}*)
thm alpha_induct
thm alpha.induct
fun
option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
where
"option_map f (nSome x) = nSome (f x)"
| "option_map f nNone = nNone"
fun
option_rel
where
"option_rel r (nSome x) (nSome y) = r x y"
| "option_rel r _ _ = False"
declare [[map noption = (option_map, option_rel)]]
lemma OPT_QUOTIENT:
assumes q: "QUOTIENT R Abs Rep"
shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)"
apply (unfold QUOTIENT_def)
apply (auto)
using q
apply (unfold QUOTIENT_def)
apply (case_tac "a :: 'b noption")
apply (simp)
apply (simp)
apply (case_tac "a :: 'b noption")
apply (simp only: option_map.simps)
apply (subst option_rel.simps)
(* Simp starts hanging so don't know how to continue *)
sorry