theory LamEx
imports Nominal QuotMain
begin
atom_decl name
nominal_datatype rlam =
rVar "name"
| rApp "rlam" "rlam"
| rLam "name" "rlam"
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\<sharp>[b].s\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
quotient lam = rlam / alpha
sorry
print_quotients
quotient_def (for lam)
Var :: "name \<Rightarrow> lam"
where
"Var \<equiv> rVar"
quotient_def (for lam)
App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
where
"App \<equiv> rApp"
quotient_def (for lam)
Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
where
"Lam \<equiv> rLam"
thm Var_def
thm App_def
thm Lam_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 (for lam)
perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
where
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
end
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 real_alpha:
assumes "t = ([(a,b)]\<bullet>s)" "a\<sharp>s"
shows "Lam a t = Lam b s"
sorry
(* Construction Site code *)
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
ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def} *}
ML {* val consts = [@{const_name "rVar"}, @{const_name "rApp"}, @{const_name "rLam"}, @{const_name "perm"}]; *}
ML {* val rty = @{typ "rlam"} *}
ML {* val qty = @{typ "lam"} *}
ML {* val rel = @{term "alpha"} *}
ML {* val rel_eqv = (#equiv_thm o hd) (quotdata_lookup @{context}) *}
ML {* val rel_refl = @{thm EQUIV_REFL} OF [rel_eqv] *}
ML {* val quot = @{thm QUOTIENT_lam} *}
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
ML {* val trans2 = @{thm QUOT_TYPE_I_lam.R_trans2} *}
ML {* val reps_same = @{thm QUOT_TYPE_I_lam.REPS_same} *}
ML {* add_lower_defs @{context} @{thms perm_lam_def} *}
ML {* val rr = @{thm eq_reflection} OF [hd (rev (add_lower_defs @{context} @{thms perm_lam_def}))] *}
lemma prod_fun_id: "prod_fun id id = id"
apply (simp add: prod_fun_def)
done
lemma map_id: "map id x = x"
apply (induct x)
apply (simp_all add: map.simps)
done
ML {* val rrr = repeat_eqsubst_thm @{context} @{thms prod_fun_id map_id} rr *}
ML {*
fun lift_thm_lam lthy t =
lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same defs t
*}
ML {* val t1 = Toplevel.program (fn () => lift_thm_lam @{context} @{thm pi_var_com}) *}
ML {* val t2 = eqsubst_thm @{context} [rrr] (atomize_thm t1) *}
ML {* ObjectLogic.rulify t2 *}
ML {* val t1 = Toplevel.program (fn () => lift_thm_lam @{context} @{thm pi_app_com}) *}
ML {* val t2 = repeat_eqsubst_thm @{context} [rrr] (atomize_thm t1) *}
ML {* ObjectLogic.rulify t2 *}
ML {* val t1 = Toplevel.program (fn () => lift_thm_lam @{context} @{thm pi_lam_com}) *}
ML {* val t2 = repeat_eqsubst_thm @{context} [rrr] (atomize_thm t1) *}
ML {* ObjectLogic.rulify t2 *}