LamEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 28 Oct 2009 18:08:38 +0100
changeset 228 268a727b0f10
parent 227 722fa927e505
child 232 38810e1df801
permissions -rw-r--r--
disambiguate ===> syntax

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>s\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"

quotient lam = rlam / alpha
apply -
sorry

print_quotients

local_setup {*
  old_make_const_def @{binding Var} @{term "rVar"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd #>
  old_make_const_def @{binding App} @{term "rApp"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd #>
  old_make_const_def @{binding Lam} @{term "rLam"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd
*}

thm Var_def
thm App_def
thm Lam_def

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>" sorry
lemma fresh_rsp: "(op = ===> (alpha ===> op =)) fresh fresh" sorry
lemma rLam_rsp: "(op = ===> (alpha ===> alpha)) rLam rLam" sorry

ML {* val defs = @{thms Var_def App_def Lam_def} *}
ML {* val consts = [@{const_name "rVar"}, @{const_name "rApp"}, @{const_name "rLam"}]; *}

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 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 {*
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 {* lift_thm_lam @{context} @{thm a3} *}

local_setup {*
  old_make_const_def @{binding lperm} @{term "perm :: ('a \<times> 'a) list \<Rightarrow> rlam \<Rightarrow> rlam"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd
*}

ML {* val consts = @{const_name perm} :: consts *}
ML {* val defs = @{thms lperm_def} @ defs *}
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 t_b = lift_thm_lam @{context} @{thm a3} *}
ML {* val rr = @{thm eq_reflection} OF [hd (rev (add_lower_defs @{context} @{thms lperm_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 {* val t_b' = eqsubst_thm @{context} [rrr] (atomize_thm t_b) *}
ML {* ObjectLogic.rulify t_b' *}

ML {* val t_u = MetaSimplifier.rewrite_rule @{thms fresh_def supp_def} @{thm a3} *}
ML {* Toplevel.program (fn () => lift_thm_lam @{context} t_u) *} (* Is not ture *)