Nominal/Manual/Term5n.thy
branchNominal2-Isabelle2012
changeset 3171 f5057aabf5c0
parent 3170 89715c48f728
equal deleted inserted replaced
3170:89715c48f728 3171:f5057aabf5c0
     1 theory Term5
       
     2 imports "../Nominal2_Atoms" "../Nominal2_Eqvt" "../Nominal2_Supp" "../Abs" "../Perm" "../Fv" "../Rsp"
       
     3 begin
       
     4 
       
     5 atom_decl name
       
     6 
       
     7 datatype rtrm5 =
       
     8   rVr5 "name"
       
     9 | rAp5 "rtrm5" "rtrm5"
       
    10 | rLt5 "rlts" "rtrm5" --"bind (bv5 lts) in (rtrm5)"
       
    11 and rlts =
       
    12   rLnil
       
    13 | rLcons "name" "rtrm5" "rlts"
       
    14 
       
    15 primrec
       
    16   rbv5
       
    17 where
       
    18   "rbv5 rLnil = {}"
       
    19 | "rbv5 (rLcons n t ltl) = {atom n} \<union> (rbv5 ltl)"
       
    20 
       
    21 
       
    22 setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Term5.rtrm5") 2 *}
       
    23 print_theorems
       
    24 
       
    25 local_setup {* snd o define_fv_alpha (Datatype.the_info @{theory} "Term5.rtrm5")
       
    26   [[[], [], [(SOME (@{term rbv5}, false), 0, 1)]], [[], []]] [(@{term rbv5}, 1, [[], [(0,NONE), (2,SOME @{term rbv5})]])] *}
       
    27 print_theorems
       
    28 
       
    29 notation
       
    30   alpha_rtrm5 ("_ \<approx>5 _" [100, 100] 100) and
       
    31   alpha_rlts ("_ \<approx>l _" [100, 100] 100)
       
    32 thm alpha_rtrm5_alpha_rlts_alpha_rbv5.intros
       
    33 
       
    34 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_rel_inj @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} ctxt)) ctxt)) *}
       
    35 thm alpha5_inj
       
    36 
       
    37 lemma rbv5_eqvt[eqvt]:
       
    38   "pi \<bullet> (rbv5 x) = rbv5 (pi \<bullet> x)"
       
    39   apply (induct x)
       
    40   apply (simp_all add: eqvts atom_eqvt)
       
    41   done
       
    42 
       
    43 lemma fv_rtrm5_rlts_eqvt[eqvt]:
       
    44   "pi \<bullet> (fv_rtrm5 x) = fv_rtrm5 (pi \<bullet> x)"
       
    45   "pi \<bullet> (fv_rlts l) = fv_rlts (pi \<bullet> l)"
       
    46   "pi \<bullet> (fv_rbv5 l) = fv_rbv5 (pi \<bullet> l)"
       
    47   apply (induct x and l)
       
    48   apply (simp_all add: eqvts atom_eqvt)
       
    49   done
       
    50 
       
    51 local_setup {*
       
    52 (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_eqvt}, []),
       
    53 build_alpha_eqvts [@{term alpha_rtrm5}, @{term alpha_rlts}, @{term alpha_rbv5}] (fn _ => alpha_eqvt_tac  @{thm alpha_rtrm5_alpha_rlts_alpha_rbv5.induct} @{thms alpha5_inj permute_rtrm5_permute_rlts.simps} ctxt 1) ctxt) ctxt)) *}
       
    54 print_theorems
       
    55 
       
    56 lemma alpha5_reflp:
       
    57 "y \<approx>5 y \<and> (x \<approx>l x \<and> alpha_rbv5 x x)"
       
    58 apply (rule rtrm5_rlts.induct)
       
    59 apply (simp_all add: alpha5_inj)
       
    60 apply (rule_tac x="0::perm" in exI)
       
    61 apply (simp add: eqvts alpha_gen fresh_star_def fresh_zero_perm)
       
    62 done
       
    63 
       
    64 lemma alpha5_symp:
       
    65 "(a \<approx>5 b \<longrightarrow> b \<approx>5 a) \<and>
       
    66 (x \<approx>l y \<longrightarrow> y \<approx>l x) \<and>
       
    67 (alpha_rbv5 x y \<longrightarrow> alpha_rbv5 y x)"
       
    68 apply (rule alpha_rtrm5_alpha_rlts_alpha_rbv5.induct)
       
    69 apply (simp_all add: alpha5_inj)
       
    70 apply (erule conjE)+
       
    71 apply (erule exE)
       
    72 apply (rule_tac x="-pi" in exI)
       
    73 apply (rule alpha_gen_sym)
       
    74 apply (simp add: alphas)
       
    75 apply (simp add: alpha5_eqvt)
       
    76 apply (simp add: alphas)
       
    77 apply clarify
       
    78 apply simp
       
    79 done
       
    80 
       
    81 lemma alpha5_transp:
       
    82 "(a \<approx>5 b \<longrightarrow> (\<forall>c. b \<approx>5 c \<longrightarrow> a \<approx>5 c)) \<and>
       
    83 (x \<approx>l y \<longrightarrow> (\<forall>z. y \<approx>l z \<longrightarrow> x \<approx>l z)) \<and>
       
    84 (alpha_rbv5 k l \<longrightarrow> (\<forall>m. alpha_rbv5 l m \<longrightarrow> alpha_rbv5 k m))"
       
    85 apply (rule alpha_rtrm5_alpha_rlts_alpha_rbv5.induct)
       
    86 apply (rule_tac [!] allI)
       
    87 apply (simp_all add: alpha5_inj)
       
    88 apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *})
       
    89 apply (simp_all add: alpha5_inj)
       
    90 defer
       
    91 apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *})
       
    92 apply (simp_all add: alpha5_inj)
       
    93 apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *})
       
    94 apply (simp_all add: alpha5_inj)
       
    95 apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *})
       
    96 apply (simp_all add: alpha5_inj)
       
    97 apply (erule conjE)+
       
    98 apply (erule exE)+
       
    99 apply (rule_tac x="pi + pia" in exI)
       
   100 apply (rule alpha_gen_trans)
       
   101 prefer 6
       
   102 apply assumption
       
   103 apply (simp_all add: alphas alpha5_eqvt)
       
   104 apply (clarify)
       
   105 apply simp
       
   106 done
       
   107 
       
   108 lemma alpha5_equivp:
       
   109   "equivp alpha_rtrm5"
       
   110   "equivp alpha_rlts"
       
   111   unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
       
   112   apply (simp_all only: alpha5_reflp)
       
   113   apply (meson alpha5_symp alpha5_transp)
       
   114   apply (meson alpha5_symp alpha5_transp)
       
   115   done
       
   116 
       
   117 quotient_type
       
   118   trm5 = rtrm5 / alpha_rtrm5
       
   119 and
       
   120   lts = rlts / alpha_rlts
       
   121   by (auto intro: alpha5_equivp)
       
   122 
       
   123 local_setup {*
       
   124 (fn ctxt => ctxt
       
   125  |> snd o (Quotient_Def.quotient_lift_const ("Vr5", @{term rVr5}))
       
   126  |> snd o (Quotient_Def.quotient_lift_const ("Ap5", @{term rAp5}))
       
   127  |> snd o (Quotient_Def.quotient_lift_const ("Lt5", @{term rLt5}))
       
   128  |> snd o (Quotient_Def.quotient_lift_const ("Lnil", @{term rLnil}))
       
   129  |> snd o (Quotient_Def.quotient_lift_const ("Lcons", @{term rLcons}))
       
   130  |> snd o (Quotient_Def.quotient_lift_const ("fv_trm5", @{term fv_rtrm5}))
       
   131  |> snd o (Quotient_Def.quotient_lift_const ("fv_lts", @{term fv_rlts}))
       
   132  |> snd o (Quotient_Def.quotient_lift_const ("fv_bv5", @{term fv_rbv5}))
       
   133  |> snd o (Quotient_Def.quotient_lift_const ("bv5", @{term rbv5}))
       
   134  |> snd o (Quotient_Def.quotient_lift_const ("alpha_bv5", @{term alpha_rbv5})))
       
   135 *}
       
   136 print_theorems
       
   137 
       
   138 lemma alpha5_rfv:
       
   139   "(t \<approx>5 s \<Longrightarrow> fv_rtrm5 t = fv_rtrm5 s)"
       
   140   "(l \<approx>l m \<Longrightarrow> (fv_rlts l = fv_rlts m \<and> fv_rbv5 l = fv_rbv5 m))"
       
   141   "(alpha_rbv5 b c \<Longrightarrow> fv_rbv5 b = fv_rbv5 c)"
       
   142   apply(induct rule: alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts)
       
   143   apply(simp_all)
       
   144   apply(simp add: alpha_gen)
       
   145   done
       
   146 
       
   147 lemma bv_list_rsp:
       
   148   shows "x \<approx>l y \<Longrightarrow> rbv5 x = rbv5 y"
       
   149   apply(induct rule: alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts(2))
       
   150   apply(simp_all)
       
   151   apply(clarify)
       
   152   apply simp
       
   153   done
       
   154 
       
   155 local_setup {* snd o Local_Theory.note ((@{binding alpha_dis}, []), (flat (map (distinct_rel @{context} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases}) [(@{thms rtrm5.distinct}, @{term alpha_rtrm5}), (@{thms rlts.distinct}, @{term alpha_rlts}), (@{thms rlts.distinct}, @{term alpha_rbv5})]))) *}
       
   156 print_theorems
       
   157 
       
   158 local_setup {* snd o Local_Theory.note ((@{binding alpha_bn_rsp}, []), prove_alpha_bn_rsp [@{term alpha_rtrm5}, @{term alpha_rlts}] @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.inducts} @{thms alpha5_inj alpha_dis} @{thms alpha5_equivp} @{context} (@{term alpha_rbv5}, 1)) *}
       
   159 thm alpha_bn_rsp
       
   160 
       
   161 
       
   162 lemma [quot_respect]:
       
   163   "(alpha_rlts ===> op =) fv_rlts fv_rlts"
       
   164   "(alpha_rlts ===> op =) fv_rbv5 fv_rbv5"
       
   165   "(alpha_rtrm5 ===> op =) fv_rtrm5 fv_rtrm5"
       
   166   "(alpha_rlts ===> op =) rbv5 rbv5"
       
   167   "(op = ===> alpha_rtrm5) rVr5 rVr5"
       
   168   "(alpha_rtrm5 ===> alpha_rtrm5 ===> alpha_rtrm5) rAp5 rAp5"
       
   169   "(alpha_rlts ===> alpha_rtrm5 ===> alpha_rtrm5) rLt5 rLt5"
       
   170   "(op = ===> alpha_rtrm5 ===> alpha_rlts ===> alpha_rlts) rLcons rLcons"
       
   171   "(op = ===> alpha_rtrm5 ===> alpha_rtrm5) permute permute"
       
   172   "(op = ===> alpha_rlts ===> alpha_rlts) permute permute"
       
   173   "(alpha_rlts ===> alpha_rlts ===> op =) alpha_rbv5 alpha_rbv5"
       
   174   apply (simp_all add: alpha5_inj alpha5_rfv alpha5_eqvt bv_list_rsp alpha5_reflp alpha_bn_rsp)
       
   175   apply (clarify)
       
   176   apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
       
   177 done
       
   178 
       
   179 lemma
       
   180   shows "(alpha_rlts ===> op =) rbv5 rbv5"
       
   181   by (simp add: bv_list_rsp)
       
   182 
       
   183 lemmas trm5_lts_inducts = rtrm5_rlts.inducts[quot_lifted]
       
   184 
       
   185 instantiation trm5 and lts :: pt
       
   186 begin
       
   187 
       
   188 quotient_definition
       
   189   "permute_trm5 :: perm \<Rightarrow> trm5 \<Rightarrow> trm5"
       
   190 is
       
   191   "permute :: perm \<Rightarrow> rtrm5 \<Rightarrow> rtrm5"
       
   192 
       
   193 quotient_definition
       
   194   "permute_lts :: perm \<Rightarrow> lts \<Rightarrow> lts"
       
   195 is
       
   196   "permute :: perm \<Rightarrow> rlts \<Rightarrow> rlts"
       
   197 
       
   198 instance by default
       
   199   (simp_all add: permute_rtrm5_permute_rlts_zero[quot_lifted] permute_rtrm5_permute_rlts_append[quot_lifted])
       
   200 
       
   201 end
       
   202 
       
   203 lemmas permute_trm5_lts = permute_rtrm5_permute_rlts.simps[quot_lifted]
       
   204 lemmas bv5[simp] = rbv5.simps[quot_lifted]
       
   205 lemmas fv_trm5_bv5[simp] = fv_rtrm5_fv_rbv5.simps[quot_lifted]
       
   206 lemmas fv_lts[simp] = fv_rlts.simps[quot_lifted]
       
   207 lemmas alpha5_INJ = alpha5_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
       
   208 
       
   209 end