LamEx.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 28 Oct 2009 15:25:11 +0100
changeset 218 df05cd030d2f
parent 203 7384115df9fd
child 219 329111e1c4ba
permissions -rw-r--r--
added infrastructure for defining lifted constants

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