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
(*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 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 qty = @{typ "lam"} *}
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def} *}
ML {* val consts = lookup_quot_consts defs *}
ML {* val (trans2, reps_same, quot) = lookup_quot_thms @{context} "lam" *}
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_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 {* lift_thm_lam @{context} @{thm pi_var_com} *}
ML {* lift_thm_lam @{context} @{thm pi_app_com} *}
ML {* lift_thm_lam @{context} @{thm pi_lam_com} *}
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
(* Not sure if it make sense or if it will be needed *)
lemma abs_fun_rsp: "(op = ===> alpha ===> op = ===> op =) abs_fun abs_fun"
sorry
(* Should not be needed *)
lemma eq_rsp2: "((op = ===> op =) ===> (op = ===> op =) ===> op =) op = op ="
apply auto
apply (rule ext)
apply auto
apply (rule ext)
apply auto
done
(* Should not be needed *)
lemma perm_rsp_eq: "(op = ===> (op = ===> op =) ===> op = ===> op =) op \<bullet> op \<bullet>"
apply auto
thm arg_cong2
apply (rule_tac f="perm x" in arg_cong2)
apply (auto)
apply (rule ext)
apply (auto)
done
(* Should not be needed *)
lemma fresh_rsp_eq: "(op = ===> (op = ===> op =) ===> op =) fresh fresh"
apply (simp add: FUN_REL.simps)
apply (metis ext)
done
(* It is just a test, it doesn't seem true... *)
lemma quotient_cheat: "QUOTIENT op = (option_map ABS_lam) (option_map REP_lam)"
sorry
ML {* val rsp_thms = @{thms abs_fun_rsp OPT_QUOTIENT eq_rsp2 quotient_cheat perm_rsp_eq fresh_rsp_eq} @ rsp_thms *}
ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
thm a3
ML {* Toplevel.program (fn () => lift_thm_lam @{context} @{thm a3}) *}
thm a3
ML {* val t_u1 = eqsubst_thm @{context} @{thms abs_fresh(1)} (atomize_thm @{thm a3}) *}
ML {* val t_u = MetaSimplifier.rewrite_rule @{thms fresh_def supp_def} t_u1 *}
ML {* val t_a = atomize_thm t_u *}
ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}
ML {* fun r_mk_comb_tac_lam ctxt =
r_mk_comb_tac ctxt rty quot rel_refl trans2 rsp_thms
*}
instance lam :: fs_name
apply(intro_classes)
sorry
prove asdf: {* Logic.mk_implies (concl_of t_r, (@{term "Trueprop (\<forall>t\<Colon>rlam\<in>Respects
alpha.
\<forall>(a\<Colon>name) b\<Colon>name.
\<forall>s\<Colon>rlam\<in>Respects
alpha.
t \<approx> ([(a,
b)] \<bullet> s) \<longrightarrow>
a = b \<or>
a
\<notin> {a\<Colon>name.
infinite
{b\<Colon>name. Not
(([(a, b)] \<bullet>
s) \<approx>
s)}} \<longrightarrow>
rLam a
t \<approx> rLam
b s)"})) *}
apply (tactic {* full_simp_tac ((Simplifier.context @{context} HOL_ss) addsimps
[(@{thm equiv_res_forall} OF [rel_eqv]),
(@{thm equiv_res_exists} OF [rel_eqv])]) 1 *})
apply (rule allI)
apply (drule_tac x="t" in spec)
apply (rule allI)
apply (drule_tac x="a" in spec)
apply (rule allI)
apply (drule_tac x="b" in spec)
apply (rule allI)
apply (drule_tac x="s" in spec)
apply (rule impI)
apply (drule_tac mp)
apply (simp)
apply (simp)
apply (rule impI)
apply (rule a3)
apply (simp)
apply (simp add: abs_fresh(1))
apply (case_tac "a = b")
apply (simp)
apply (simp)
apply (auto)
apply (unfold fresh_def)
apply (unfold supp_def)
apply (simp)
prefer 2
apply (simp)
sorry
ML {* val abs = findabs rty (prop_of t_a) *}
ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *}
ML {* val t_defs_sym = add_lower_defs @{context} defs *}
ML {* val t_r' = @{thm asdf} OF [t_r] *}
ML {* val t_t = repabs @{context} t_r' consts rty qty quot rel_refl trans2 rsp_thms *}
ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
ML {* val t_a = simp_allex_prs @{context} quot t_l *}
ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_a *}
ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
ML {* val tt = MetaSimplifier.rewrite_rule [symmetric @{thm supp_def}, symmetric @{thm fresh_def}] t_r *}
ML {* val rr = @{thm sym} OF @{thms abs_fresh(1)} *}
ML {* val ttt = eqsubst_thm @{context} [rr] tt *}
ML {* ObjectLogic.rulify ttt *}
lemma
assumes a: "a \<notin> {a\<Colon>name. infinite {b\<Colon>name. \<not> ([(a, b)] \<bullet> s) \<approx> s}}"
shows "a \<notin> {a\<Colon>name. infinite {b\<Colon>name. [(a, b)] \<bullet> s \<noteq> s}}"
using a apply simp
sorry (* Not true... *)