LamEx.thy
changeset 597 8a1c8dc72b5c
parent 596 6088fea1c8b1
child 598 ae254a6d685c
equal deleted inserted replaced
596:6088fea1c8b1 597:8a1c8dc72b5c
     1 theory LamEx
       
     2 imports Nominal QuotMain
       
     3 begin
       
     4 
       
     5 atom_decl name
       
     6 
       
     7 thm abs_fresh(1)
       
     8 
       
     9 nominal_datatype rlam =
       
    10   rVar "name"
       
    11 | rApp "rlam" "rlam"
       
    12 | rLam "name" "rlam"
       
    13 
       
    14 print_theorems
       
    15 
       
    16 function
       
    17   rfv :: "rlam \<Rightarrow> name set"
       
    18 where
       
    19   rfv_var: "rfv (rVar a) = {a}"
       
    20 | rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
       
    21 | rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
       
    22 sorry
       
    23 
       
    24 termination rfv sorry
       
    25 
       
    26 inductive
       
    27   alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
       
    28 where
       
    29   a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
       
    30 | a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
       
    31 | a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
       
    32 
       
    33 print_theorems
       
    34 
       
    35 lemma alpha_refl:
       
    36   fixes t::"rlam"
       
    37   shows "t \<approx> t"
       
    38   apply(induct t rule: rlam.induct)
       
    39   apply(simp add: a1)
       
    40   apply(simp add: a2)
       
    41   apply(rule a3)
       
    42   apply(subst pt_swap_bij'')
       
    43   apply(rule pt_name_inst)
       
    44   apply(rule at_name_inst)
       
    45   apply(simp)
       
    46   apply(simp)
       
    47   done
       
    48 
       
    49 lemma alpha_equivp:
       
    50   shows "equivp alpha"
       
    51 sorry
       
    52 
       
    53 quotient lam = rlam / alpha
       
    54   apply(rule alpha_equivp)
       
    55   done
       
    56 
       
    57 print_quotients
       
    58 
       
    59 quotient_def 
       
    60   Var :: "name \<Rightarrow> lam"
       
    61 where
       
    62   "Var \<equiv> rVar"
       
    63 
       
    64 quotient_def 
       
    65   App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
       
    66 where
       
    67   "App \<equiv> rApp"
       
    68 
       
    69 quotient_def 
       
    70   Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
       
    71 where
       
    72   "Lam \<equiv> rLam"
       
    73 
       
    74 thm Var_def
       
    75 thm App_def
       
    76 thm Lam_def
       
    77 
       
    78 quotient_def 
       
    79   fv :: "lam \<Rightarrow> name set"
       
    80 where
       
    81   "fv \<equiv> rfv"
       
    82 
       
    83 thm fv_def
       
    84 
       
    85 (* definition of overloaded permutation function *)
       
    86 (* for the lifted type lam                       *)
       
    87 overloading
       
    88   perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"   (unchecked)
       
    89 begin
       
    90 
       
    91 quotient_def 
       
    92   perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
       
    93 where
       
    94   "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
       
    95 
       
    96 end
       
    97 
       
    98 (*quotient_def (for lam)
       
    99   abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
       
   100 where
       
   101   "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*)
       
   102 
       
   103 
       
   104 thm perm_lam_def
       
   105 
       
   106 (* lemmas that need to lift *)
       
   107 lemma pi_var_com:
       
   108   fixes pi::"'x prm"
       
   109   shows "(pi\<bullet>rVar a) \<approx> rVar (pi\<bullet>a)"
       
   110   sorry
       
   111 
       
   112 lemma pi_app_com:
       
   113   fixes pi::"'x prm"
       
   114   shows "(pi\<bullet>rApp t1 t2) \<approx> rApp (pi\<bullet>t1) (pi\<bullet>t2)"
       
   115   sorry
       
   116 
       
   117 lemma pi_lam_com:
       
   118   fixes pi::"'x prm"
       
   119   shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)"
       
   120   sorry
       
   121 
       
   122 
       
   123 
       
   124 lemma real_alpha:
       
   125   assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
       
   126   shows "Lam a t = Lam b s"
       
   127 using a
       
   128 unfolding fresh_def supp_def
       
   129 sorry
       
   130 
       
   131 lemma perm_rsp[quotient_rsp]:
       
   132   "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
       
   133   apply(auto)
       
   134   (* this is propably true if some type conditions are imposed ;o) *)
       
   135   sorry
       
   136 
       
   137 lemma fresh_rsp:
       
   138   "(op = ===> alpha ===> op =) fresh fresh"
       
   139   apply(auto)
       
   140   (* this is probably only true if some type conditions are imposed *)
       
   141   sorry
       
   142 
       
   143 lemma rVar_rsp[quotient_rsp]:
       
   144   "(op = ===> alpha) rVar rVar"
       
   145 by (auto intro:a1)
       
   146 
       
   147 lemma rApp_rsp[quotient_rsp]: "(alpha ===> alpha ===> alpha) rApp rApp"
       
   148 by (auto intro:a2)
       
   149 
       
   150 lemma rLam_rsp[quotient_rsp]: "(op = ===> alpha ===> alpha) rLam rLam"
       
   151   apply(auto)
       
   152   apply(rule a3)
       
   153   apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst)
       
   154   apply(rule sym)
       
   155   apply(rule trans)
       
   156   apply(rule pt_name3)
       
   157   apply(rule at_ds1[OF at_name_inst])
       
   158   apply(simp add: pt_name1)
       
   159   apply(assumption)
       
   160   apply(simp add: abs_fresh)
       
   161   done
       
   162 
       
   163 lemma rfv_rsp[quotient_rsp]: "(alpha ===> op =) rfv rfv"
       
   164   sorry
       
   165 
       
   166 lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)"
       
   167 apply (auto)
       
   168 apply (erule alpha.cases)
       
   169 apply (simp_all add: rlam.inject alpha_refl)
       
   170 done
       
   171 
       
   172 ML {* val qty = @{typ "lam"} *}
       
   173 ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *}
       
   174 
       
   175 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
       
   176 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
       
   177 ML {* fun lift_tac_lam lthy t = lift_tac lthy t *}
       
   178 
       
   179 lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
       
   180 apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *})
       
   181 done
       
   182 
       
   183 lemma pi_app: "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
       
   184 apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *})
       
   185 done
       
   186 
       
   187 lemma pi_lam: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
       
   188 apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *})
       
   189 done
       
   190 
       
   191 lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
       
   192 apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *})
       
   193 done
       
   194 
       
   195 lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa"
       
   196 apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *})
       
   197 done
       
   198 
       
   199 lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
       
   200 apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *})
       
   201 done
       
   202 
       
   203 lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
       
   204 apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *})
       
   205 done
       
   206 
       
   207 lemma a2: "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
       
   208 apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *})
       
   209 done
       
   210 
       
   211 lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa"
       
   212 apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *})
       
   213 done
       
   214 
       
   215 lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
       
   216      \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
       
   217      \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
       
   218     \<Longrightarrow> P"
       
   219 apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *})
       
   220 done
       
   221 
       
   222 lemma alpha_induct: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b);
       
   223      \<And>(x\<Colon>lam) (xa\<Colon>lam) (xb\<Colon>lam) xc\<Colon>lam. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
       
   224      \<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam.
       
   225         \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
       
   226     \<Longrightarrow> qxb qx qxa"
       
   227 apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *})
       
   228 done
       
   229 
       
   230 lemma var_inject: "(Var a = Var b) = (a = b)"
       
   231 apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *})
       
   232 done
       
   233 
       
   234 lemma lam_induct:" \<lbrakk>\<And>name. P (Var name); \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
       
   235               \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> \<Longrightarrow> P lam"
       
   236 apply (tactic {* lift_tac_lam @{context} @{thm rlam.induct} 1 *})
       
   237 done
       
   238 
       
   239 lemma var_supp:
       
   240   shows "supp (Var a) = ((supp a)::name set)"
       
   241   apply(simp add: supp_def)
       
   242   apply(simp add: pi_var)
       
   243   apply(simp add: var_inject)
       
   244   done
       
   245 
       
   246 lemma var_fresh:
       
   247   fixes a::"name"
       
   248   shows "(a\<sharp>(Var b)) = (a\<sharp>b)"
       
   249   apply(simp add: fresh_def)
       
   250   apply(simp add: var_supp)
       
   251   done
       
   252