src/HOL/Nominal/Examples/Class2.thy
changeset 36277 9be4ab2acc13
child 41893 dde7df1176b7
equal deleted inserted replaced
36276:92011cc923f5 36277:9be4ab2acc13
       
     1 theory Class2
       
     2 imports Class1
       
     3 begin
       
     4 
       
     5 text {* Reduction *}
       
     6 
       
     7 lemma fin_not_Cut:
       
     8   assumes a: "fin M x"
       
     9   shows "\<not>(\<exists>a M' x N'. M = Cut <a>.M' (x).N')"
       
    10 using a
       
    11 by (induct) (auto)
       
    12 
       
    13 lemma fresh_not_fin:
       
    14   assumes a: "x\<sharp>M"
       
    15   shows "\<not>fin M x"
       
    16 proof -
       
    17   have "fin M x \<Longrightarrow> x\<sharp>M \<Longrightarrow> False" by (induct rule: fin.induct) (auto simp add: abs_fresh fresh_atm)
       
    18   with a show "\<not>fin M x" by blast
       
    19 qed
       
    20 
       
    21 lemma fresh_not_fic:
       
    22   assumes a: "a\<sharp>M"
       
    23   shows "\<not>fic M a"
       
    24 proof -
       
    25   have "fic M a \<Longrightarrow> a\<sharp>M \<Longrightarrow> False" by (induct rule: fic.induct) (auto simp add: abs_fresh fresh_atm)
       
    26   with a show "\<not>fic M a" by blast
       
    27 qed
       
    28 
       
    29 lemma c_redu_subst1:
       
    30   assumes a: "M \<longrightarrow>\<^isub>c M'" "c\<sharp>M" "y\<sharp>P"
       
    31   shows "M{y:=<c>.P} \<longrightarrow>\<^isub>c M'{y:=<c>.P}"
       
    32 using a
       
    33 proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct)
       
    34   case (left M a N x)
       
    35   then show ?case
       
    36     apply -
       
    37     apply(simp)
       
    38     apply(rule conjI)
       
    39     apply(force)
       
    40     apply(auto)
       
    41     apply(subgoal_tac "M{a:=(x).N}{y:=<c>.P} = M{y:=<c>.P}{a:=(x).(N{y:=<c>.P})}")(*A*)
       
    42     apply(simp)
       
    43     apply(rule c_redu.intros)
       
    44     apply(rule not_fic_subst1)
       
    45     apply(simp)
       
    46     apply(simp add: subst_fresh)
       
    47     apply(simp add: subst_fresh)
       
    48     apply(simp add: abs_fresh fresh_atm)
       
    49     apply(rule subst_subst2)
       
    50     apply(simp add: fresh_prod fresh_atm)
       
    51     apply(simp add: fresh_prod fresh_atm)
       
    52     apply(simp add: fresh_prod fresh_atm)
       
    53     apply(simp)
       
    54     done
       
    55 next
       
    56   case (right N x a M)
       
    57   then show ?case
       
    58     apply -
       
    59     apply(simp)
       
    60     apply(rule conjI)
       
    61     (* case M = Ax y a *)
       
    62     apply(rule impI)
       
    63     apply(subgoal_tac "N{x:=<a>.Ax y a}{y:=<c>.P} = N{y:=<c>.P}{x:=<c>.P}")
       
    64     apply(simp)
       
    65     apply(rule c_redu.right)
       
    66     apply(rule not_fin_subst2)
       
    67     apply(simp)
       
    68     apply(rule subst_fresh)
       
    69     apply(simp add: abs_fresh)
       
    70     apply(simp add: abs_fresh)
       
    71     apply(rule sym)
       
    72     apply(rule interesting_subst1')
       
    73     apply(simp add: fresh_atm)
       
    74     apply(simp)
       
    75     apply(simp)
       
    76     (* case M \<noteq> Ax y a*)
       
    77     apply(rule impI)
       
    78     apply(subgoal_tac "N{x:=<a>.M}{y:=<c>.P} = N{y:=<c>.P}{x:=<a>.(M{y:=<c>.P})}")
       
    79     apply(simp)
       
    80     apply(rule c_redu.right)
       
    81     apply(rule not_fin_subst2)
       
    82     apply(simp)
       
    83     apply(simp add: subst_fresh)
       
    84     apply(simp add: subst_fresh)
       
    85     apply(simp add: abs_fresh fresh_atm)
       
    86     apply(rule subst_subst3)
       
    87     apply(simp_all add: fresh_atm fresh_prod)
       
    88     done
       
    89 qed
       
    90 
       
    91 lemma c_redu_subst2:
       
    92   assumes a: "M \<longrightarrow>\<^isub>c M'" "c\<sharp>P" "y\<sharp>M"
       
    93   shows "M{c:=(y).P} \<longrightarrow>\<^isub>c M'{c:=(y).P}"
       
    94 using a
       
    95 proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct)
       
    96   case (right N x a M)
       
    97   then show ?case
       
    98     apply -
       
    99     apply(simp)
       
   100     apply(rule conjI)
       
   101     apply(force)
       
   102     apply(auto)
       
   103     apply(subgoal_tac "N{x:=<a>.M}{c:=(y).P} = N{c:=(y).P}{x:=<a>.(M{c:=(y).P})}")(*A*)
       
   104     apply(simp)
       
   105     apply(rule c_redu.intros)
       
   106     apply(rule not_fin_subst1)
       
   107     apply(simp)
       
   108     apply(simp add: subst_fresh)
       
   109     apply(simp add: subst_fresh)
       
   110     apply(simp add: abs_fresh fresh_atm)
       
   111     apply(rule subst_subst1)
       
   112     apply(simp add: fresh_prod fresh_atm)
       
   113     apply(simp add: fresh_prod fresh_atm)
       
   114     apply(simp add: fresh_prod fresh_atm)
       
   115     apply(simp)
       
   116     done
       
   117 next
       
   118   case (left M a N x)
       
   119   then show ?case
       
   120     apply -
       
   121     apply(simp)
       
   122     apply(rule conjI)
       
   123     (* case N = Ax x c *)
       
   124     apply(rule impI)
       
   125     apply(subgoal_tac "M{a:=(x).Ax x c}{c:=(y).P} = M{c:=(y).P}{a:=(y).P}")
       
   126     apply(simp)
       
   127     apply(rule c_redu.left)
       
   128     apply(rule not_fic_subst2)
       
   129     apply(simp)
       
   130     apply(simp)
       
   131     apply(rule subst_fresh)
       
   132     apply(simp add: abs_fresh)
       
   133     apply(rule sym)
       
   134     apply(rule interesting_subst2')
       
   135     apply(simp add: fresh_atm)
       
   136     apply(simp)
       
   137     apply(simp)
       
   138     (* case M \<noteq> Ax y a*)
       
   139     apply(rule impI)
       
   140     apply(subgoal_tac "M{a:=(x).N}{c:=(y).P} = M{c:=(y).P}{a:=(x).(N{c:=(y).P})}")
       
   141     apply(simp)
       
   142     apply(rule c_redu.left)
       
   143     apply(rule not_fic_subst2)
       
   144     apply(simp)
       
   145     apply(simp add: subst_fresh)
       
   146     apply(simp add: subst_fresh)
       
   147     apply(simp add: abs_fresh fresh_atm)
       
   148     apply(rule subst_subst4)
       
   149     apply(simp add: fresh_prod fresh_atm)
       
   150     apply(simp add: fresh_prod fresh_atm)
       
   151     apply(simp add: fresh_prod fresh_atm)
       
   152     apply(simp add: fresh_prod fresh_atm)
       
   153     apply(simp)
       
   154     done
       
   155 qed
       
   156 
       
   157 lemma c_redu_subst1':
       
   158   assumes a: "M \<longrightarrow>\<^isub>c M'" 
       
   159   shows "M{y:=<c>.P} \<longrightarrow>\<^isub>c M'{y:=<c>.P}"
       
   160 using a
       
   161 proof -
       
   162   obtain y'::"name"   where fs1: "y'\<sharp>(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast)
       
   163   obtain c'::"coname" where fs2: "c'\<sharp>(M,M',P,P,c)" by (rule exists_fresh(2), rule fin_supp, blast)
       
   164   have "M{y:=<c>.P} = ([(y',y)]\<bullet>M){y':=<c'>.([(c',c)]\<bullet>P)}" using fs1 fs2
       
   165     apply -
       
   166     apply(rule trans)
       
   167     apply(rule_tac y="y'" in subst_rename(3))
       
   168     apply(simp)
       
   169     apply(rule subst_rename(4))
       
   170     apply(simp)
       
   171     done
       
   172   also have "\<dots> \<longrightarrow>\<^isub>c ([(y',y)]\<bullet>M'){y':=<c'>.([(c',c)]\<bullet>P)}" using fs1 fs2
       
   173     apply -
       
   174     apply(rule c_redu_subst1)
       
   175     apply(simp add: c_redu.eqvt a)
       
   176     apply(simp_all add: fresh_left calc_atm fresh_prod)
       
   177     done
       
   178   also have "\<dots> = M'{y:=<c>.P}" using fs1 fs2
       
   179     apply -
       
   180     apply(rule sym)
       
   181     apply(rule trans)
       
   182     apply(rule_tac y="y'" in subst_rename(3))
       
   183     apply(simp)
       
   184     apply(rule subst_rename(4))
       
   185     apply(simp)
       
   186     done
       
   187   finally show ?thesis by simp
       
   188 qed
       
   189 
       
   190 lemma c_redu_subst2':
       
   191   assumes a: "M \<longrightarrow>\<^isub>c M'" 
       
   192   shows "M{c:=(y).P} \<longrightarrow>\<^isub>c M'{c:=(y).P}"
       
   193 using a
       
   194 proof -
       
   195   obtain y'::"name"   where fs1: "y'\<sharp>(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast)
       
   196   obtain c'::"coname" where fs2: "c'\<sharp>(M,M',P,P,c)" by (rule exists_fresh(2), rule fin_supp, blast)
       
   197   have "M{c:=(y).P} = ([(c',c)]\<bullet>M){c':=(y').([(y',y)]\<bullet>P)}" using fs1 fs2
       
   198     apply -
       
   199     apply(rule trans)
       
   200     apply(rule_tac c="c'" in subst_rename(1))
       
   201     apply(simp)
       
   202     apply(rule subst_rename(2))
       
   203     apply(simp)
       
   204     done
       
   205   also have "\<dots> \<longrightarrow>\<^isub>c ([(c',c)]\<bullet>M'){c':=(y').([(y',y)]\<bullet>P)}" using fs1 fs2
       
   206     apply -
       
   207     apply(rule c_redu_subst2)
       
   208     apply(simp add: c_redu.eqvt a)
       
   209     apply(simp_all add: fresh_left calc_atm fresh_prod)
       
   210     done
       
   211   also have "\<dots> = M'{c:=(y).P}" using fs1 fs2
       
   212     apply -
       
   213     apply(rule sym)
       
   214     apply(rule trans)
       
   215     apply(rule_tac c="c'" in subst_rename(1))
       
   216     apply(simp)
       
   217     apply(rule subst_rename(2))
       
   218     apply(simp)
       
   219     done
       
   220 
       
   221   finally show ?thesis by simp
       
   222 qed
       
   223 
       
   224 lemma aux1:
       
   225   assumes a: "M = M'" "M' \<longrightarrow>\<^isub>l M''"
       
   226   shows "M \<longrightarrow>\<^isub>l M''"
       
   227 using a by simp
       
   228   
       
   229 lemma aux2:
       
   230   assumes a: "M \<longrightarrow>\<^isub>l M'" "M' = M''"
       
   231   shows "M \<longrightarrow>\<^isub>l M''"
       
   232 using a by simp
       
   233 
       
   234 lemma aux3:
       
   235   assumes a: "M = M'" "M' \<longrightarrow>\<^isub>a* M''"
       
   236   shows "M \<longrightarrow>\<^isub>a* M''"
       
   237 using a by simp
       
   238 
       
   239 lemma aux4:
       
   240   assumes a: "M = M'"
       
   241   shows "M \<longrightarrow>\<^isub>a* M'"
       
   242 using a by blast
       
   243 
       
   244 lemma l_redu_subst1:
       
   245   assumes a: "M \<longrightarrow>\<^isub>l M'" 
       
   246   shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
       
   247 using a
       
   248 proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct)
       
   249   case LAxR
       
   250   then show ?case
       
   251     apply -
       
   252     apply(rule aux3)
       
   253     apply(rule better_Cut_substn)
       
   254     apply(simp add: abs_fresh)
       
   255     apply(simp)
       
   256     apply(simp add: fresh_atm)
       
   257     apply(auto)
       
   258     apply(rule aux4)
       
   259     apply(simp add: trm.inject alpha calc_atm fresh_atm)
       
   260     apply(rule a_star_trans)
       
   261     apply(rule a_starI)
       
   262     apply(rule al_redu)
       
   263     apply(rule l_redu.intros)
       
   264     apply(simp add: subst_fresh)
       
   265     apply(simp add: fresh_atm)
       
   266     apply(rule fic_subst2)
       
   267     apply(simp_all)
       
   268     apply(rule aux4)
       
   269     apply(rule subst_comm')
       
   270     apply(simp_all)
       
   271     done
       
   272 next
       
   273   case LAxL
       
   274   then show ?case
       
   275     apply -
       
   276     apply(rule aux3)
       
   277     apply(rule better_Cut_substn)
       
   278     apply(simp add: abs_fresh)
       
   279     apply(simp)
       
   280     apply(simp add: trm.inject fresh_atm)
       
   281     apply(auto)
       
   282     apply(rule aux4)
       
   283     apply(rule sym)
       
   284     apply(rule fin_substn_nrename)
       
   285     apply(simp_all)
       
   286     apply(rule a_starI)
       
   287     apply(rule al_redu)
       
   288     apply(rule aux2)
       
   289     apply(rule l_redu.intros)
       
   290     apply(simp add: subst_fresh)
       
   291     apply(simp add: fresh_atm)
       
   292     apply(rule fin_subst1)
       
   293     apply(simp_all)
       
   294     apply(rule subst_comm')
       
   295     apply(simp_all)
       
   296     done
       
   297 next
       
   298   case (LNot v M N u a b)
       
   299   then show ?case
       
   300   proof -
       
   301     { assume asm: "N\<noteq>Ax y b"
       
   302       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} = 
       
   303         (Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using prems
       
   304         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   305       also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using prems
       
   306         by (auto intro: l_redu.intros simp add: subst_fresh)
       
   307       also have "\<dots> = (Cut <b>.N (u).M){y:=<c>.P}" using prems 
       
   308         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   309       finally have ?thesis by auto
       
   310     }
       
   311     moreover
       
   312     { assume asm: "N=Ax y b"
       
   313       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} = 
       
   314         (Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using prems
       
   315         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   316       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using prems
       
   317         apply -
       
   318         apply(rule a_starI)
       
   319         apply(rule al_redu)
       
   320         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   321         done
       
   322       also have "\<dots> = (Cut <b>.(Cut <c>.P (y).Ax y b) (u).(M{y:=<c>.P}))" using prems
       
   323         by simp
       
   324       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(P[c\<turnstile>c>b]) (u).(M{y:=<c>.P}))" 
       
   325       proof (cases "fic P c")
       
   326         case True 
       
   327         assume "fic P c"
       
   328         then show ?thesis using prems
       
   329           apply -
       
   330           apply(rule a_starI)
       
   331           apply(rule better_CutL_intro)
       
   332           apply(rule al_redu)
       
   333           apply(rule better_LAxR_intro)
       
   334           apply(simp)
       
   335           done
       
   336       next
       
   337         case False 
       
   338         assume "\<not>fic P c" 
       
   339         then show ?thesis
       
   340           apply -
       
   341           apply(rule a_star_CutL)
       
   342           apply(rule a_star_trans)
       
   343           apply(rule a_starI)
       
   344           apply(rule ac_redu)
       
   345           apply(rule better_left)
       
   346           apply(simp)
       
   347           apply(simp add: subst_with_ax2)
       
   348           done
       
   349       qed
       
   350       also have "\<dots> = (Cut <b>.N (u).M){y:=<c>.P}" using prems
       
   351         apply -
       
   352         apply(auto simp add: subst_fresh abs_fresh)
       
   353         apply(simp add: trm.inject)
       
   354         apply(simp add: alpha fresh_atm)
       
   355         apply(rule sym)
       
   356         apply(rule crename_swap)
       
   357         apply(simp)
       
   358         done
       
   359       finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){y:=<c>.P}" 
       
   360         by simp
       
   361     }
       
   362     ultimately show ?thesis by blast
       
   363   qed
       
   364 next
       
   365   case (LAnd1 b a1 M1 a2 M2 N z u)
       
   366   then show ?case
       
   367   proof -
       
   368     { assume asm: "M1\<noteq>Ax y a1"
       
   369       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} = 
       
   370         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z" 
       
   371         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   372       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
       
   373         using prems
       
   374         apply -
       
   375         apply(rule a_starI)
       
   376         apply(rule al_redu)
       
   377         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   378         done
       
   379       also have "\<dots> = (Cut <a1>.M1 (u).N){y:=<c>.P}" using prems 
       
   380         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   381       finally 
       
   382       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
       
   383         by simp
       
   384     } 
       
   385     moreover
       
   386     { assume asm: "M1=Ax y a1"
       
   387       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} = 
       
   388         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z" 
       
   389         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   390       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
       
   391         using prems
       
   392         apply -
       
   393         apply(rule a_starI)
       
   394         apply(rule al_redu)
       
   395         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   396         done
       
   397       also have "\<dots> = Cut <a1>.(Cut <c>.P (y). Ax y a1) (u).(N{y:=<c>.P})" 
       
   398         using prems by simp
       
   399       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.P[c\<turnstile>c>a1] (u).(N{y:=<c>.P})"
       
   400       proof (cases "fic P c")
       
   401         case True 
       
   402         assume "fic P c"
       
   403         then show ?thesis using prems
       
   404           apply -
       
   405           apply(rule a_starI)
       
   406           apply(rule better_CutL_intro)
       
   407           apply(rule al_redu)
       
   408           apply(rule better_LAxR_intro)
       
   409           apply(simp)
       
   410           done
       
   411       next
       
   412         case False 
       
   413         assume "\<not>fic P c" 
       
   414         then show ?thesis
       
   415           apply -
       
   416           apply(rule a_star_CutL)
       
   417           apply(rule a_star_trans)
       
   418           apply(rule a_starI)
       
   419           apply(rule ac_redu)
       
   420           apply(rule better_left)
       
   421           apply(simp)
       
   422           apply(simp add: subst_with_ax2)
       
   423           done
       
   424       qed
       
   425       also have "\<dots> = (Cut <a1>.M1 (u).N){y:=<c>.P}" using prems
       
   426         apply -
       
   427         apply(auto simp add: subst_fresh abs_fresh)
       
   428         apply(simp add: trm.inject)
       
   429         apply(simp add: alpha fresh_atm)
       
   430         apply(rule sym)
       
   431         apply(rule crename_swap)
       
   432         apply(simp)
       
   433         done
       
   434       finally 
       
   435       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
       
   436         by simp
       
   437     }
       
   438     ultimately show ?thesis by blast
       
   439   qed
       
   440 next
       
   441   case (LAnd2 b a1 M1 a2 M2 N z u)
       
   442   then show ?case
       
   443   proof -
       
   444     { assume asm: "M2\<noteq>Ax y a2"
       
   445       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} = 
       
   446         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z" 
       
   447         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   448       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
       
   449         using prems
       
   450         apply -
       
   451         apply(rule a_starI)
       
   452         apply(rule al_redu)
       
   453         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   454         done
       
   455       also have "\<dots> = (Cut <a2>.M2 (u).N){y:=<c>.P}" using prems 
       
   456         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   457       finally 
       
   458       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
       
   459         by simp
       
   460     } 
       
   461     moreover
       
   462     { assume asm: "M2=Ax y a2"
       
   463       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} = 
       
   464         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z" 
       
   465         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   466       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
       
   467         using prems
       
   468         apply -
       
   469         apply(rule a_starI)
       
   470         apply(rule al_redu)
       
   471         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   472         done
       
   473       also have "\<dots> = Cut <a2>.(Cut <c>.P (y). Ax y a2) (u).(N{y:=<c>.P})" 
       
   474         using prems by simp
       
   475       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.P[c\<turnstile>c>a2] (u).(N{y:=<c>.P})"
       
   476       proof (cases "fic P c")
       
   477         case True 
       
   478         assume "fic P c"
       
   479         then show ?thesis using prems
       
   480           apply -
       
   481           apply(rule a_starI)
       
   482           apply(rule better_CutL_intro)
       
   483           apply(rule al_redu)
       
   484           apply(rule better_LAxR_intro)
       
   485           apply(simp)
       
   486           done
       
   487       next
       
   488         case False 
       
   489         assume "\<not>fic P c" 
       
   490         then show ?thesis
       
   491           apply -
       
   492           apply(rule a_star_CutL)
       
   493           apply(rule a_star_trans)
       
   494           apply(rule a_starI)
       
   495           apply(rule ac_redu)
       
   496           apply(rule better_left)
       
   497           apply(simp)
       
   498           apply(simp add: subst_with_ax2)
       
   499           done
       
   500       qed
       
   501       also have "\<dots> = (Cut <a2>.M2 (u).N){y:=<c>.P}" using prems
       
   502         apply -
       
   503         apply(auto simp add: subst_fresh abs_fresh)
       
   504         apply(simp add: trm.inject)
       
   505         apply(simp add: alpha fresh_atm)
       
   506         apply(rule sym)
       
   507         apply(rule crename_swap)
       
   508         apply(simp)
       
   509         done
       
   510       finally 
       
   511       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
       
   512         by simp
       
   513     }
       
   514     ultimately show ?thesis by blast
       
   515   qed
       
   516 next
       
   517   case (LOr1 b a M N1 N2 z x1 x2 y c P)
       
   518   then show ?case
       
   519   proof -
       
   520     { assume asm: "M\<noteq>Ax y a"
       
   521       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
       
   522         Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
       
   523         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   524       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
       
   525         using prems
       
   526         apply -
       
   527         apply(rule a_starI)
       
   528         apply(rule al_redu)
       
   529         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   530         done
       
   531       also have "\<dots> = (Cut <a>.M (x1).N1){y:=<c>.P}" using prems 
       
   532         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   533       finally 
       
   534       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
       
   535         by simp
       
   536     } 
       
   537     moreover
       
   538     { assume asm: "M=Ax y a"
       
   539       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
       
   540         Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
       
   541         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   542       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
       
   543         using prems
       
   544         apply -
       
   545         apply(rule a_starI)
       
   546         apply(rule al_redu)
       
   547         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   548         done
       
   549       also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x1).(N1{y:=<c>.P})" 
       
   550         using prems by simp
       
   551       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x1).(N1{y:=<c>.P})"
       
   552       proof (cases "fic P c")
       
   553         case True 
       
   554         assume "fic P c"
       
   555         then show ?thesis using prems
       
   556           apply -
       
   557           apply(rule a_starI)
       
   558           apply(rule better_CutL_intro)
       
   559           apply(rule al_redu)
       
   560           apply(rule better_LAxR_intro)
       
   561           apply(simp)
       
   562           done
       
   563       next
       
   564         case False 
       
   565         assume "\<not>fic P c" 
       
   566         then show ?thesis
       
   567           apply -
       
   568           apply(rule a_star_CutL)
       
   569           apply(rule a_star_trans)
       
   570           apply(rule a_starI)
       
   571           apply(rule ac_redu)
       
   572           apply(rule better_left)
       
   573           apply(simp)
       
   574           apply(simp add: subst_with_ax2)
       
   575           done
       
   576       qed
       
   577       also have "\<dots> = (Cut <a>.M (x1).N1){y:=<c>.P}" using prems
       
   578         apply -
       
   579         apply(auto simp add: subst_fresh abs_fresh)
       
   580         apply(simp add: trm.inject)
       
   581         apply(simp add: alpha fresh_atm)
       
   582         apply(rule sym)
       
   583         apply(rule crename_swap)
       
   584         apply(simp)
       
   585         done
       
   586       finally 
       
   587       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
       
   588         by simp
       
   589     }
       
   590     ultimately show ?thesis by blast
       
   591   qed
       
   592 next
       
   593   case (LOr2 b a M N1 N2 z x1 x2 y c P)
       
   594   then show ?case
       
   595   proof -
       
   596     { assume asm: "M\<noteq>Ax y a"
       
   597       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
       
   598         Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
       
   599         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   600       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
       
   601         using prems
       
   602         apply -
       
   603         apply(rule a_starI)
       
   604         apply(rule al_redu)
       
   605         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   606         done
       
   607       also have "\<dots> = (Cut <a>.M (x2).N2){y:=<c>.P}" using prems 
       
   608         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   609       finally 
       
   610       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
       
   611         by simp
       
   612     } 
       
   613     moreover
       
   614     { assume asm: "M=Ax y a"
       
   615       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
       
   616         Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
       
   617         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   618       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
       
   619         using prems
       
   620         apply -
       
   621         apply(rule a_starI)
       
   622         apply(rule al_redu)
       
   623         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   624         done
       
   625       also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x2).(N2{y:=<c>.P})" 
       
   626         using prems by simp
       
   627       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x2).(N2{y:=<c>.P})"
       
   628       proof (cases "fic P c")
       
   629         case True 
       
   630         assume "fic P c"
       
   631         then show ?thesis using prems
       
   632           apply -
       
   633           apply(rule a_starI)
       
   634           apply(rule better_CutL_intro)
       
   635           apply(rule al_redu)
       
   636           apply(rule better_LAxR_intro)
       
   637           apply(simp)
       
   638           done
       
   639       next
       
   640         case False 
       
   641         assume "\<not>fic P c" 
       
   642         then show ?thesis
       
   643           apply -
       
   644           apply(rule a_star_CutL)
       
   645           apply(rule a_star_trans)
       
   646           apply(rule a_starI)
       
   647           apply(rule ac_redu)
       
   648           apply(rule better_left)
       
   649           apply(simp)
       
   650           apply(simp add: subst_with_ax2)
       
   651           done
       
   652       qed
       
   653       also have "\<dots> = (Cut <a>.M (x2).N2){y:=<c>.P}" using prems
       
   654         apply -
       
   655         apply(auto simp add: subst_fresh abs_fresh)
       
   656         apply(simp add: trm.inject)
       
   657         apply(simp add: alpha fresh_atm)
       
   658         apply(rule sym)
       
   659         apply(rule crename_swap)
       
   660         apply(simp)
       
   661         done
       
   662       finally 
       
   663       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
       
   664         by simp
       
   665     }
       
   666     ultimately show ?thesis by blast
       
   667   qed
       
   668 next
       
   669   case (LImp z N u Q x M b a d y c P)
       
   670   then show ?case
       
   671   proof -
       
   672     { assume asm: "N\<noteq>Ax y d"
       
   673       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} = 
       
   674         Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z" 
       
   675         using prems by (simp add: fresh_prod abs_fresh fresh_atm)
       
   676       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P})  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
       
   677         using prems
       
   678         apply -
       
   679         apply(rule a_starI)
       
   680         apply(rule al_redu)
       
   681         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   682         done
       
   683       also have "\<dots> = (Cut <a>.(Cut <d>.N  (x).M) (u).Q){y:=<c>.P}" using prems 
       
   684         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   685       finally 
       
   686       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a* 
       
   687                      (Cut <a>.(Cut <d>.N  (x).M) (u).Q){y:=<c>.P}"
       
   688         by simp
       
   689     } 
       
   690     moreover
       
   691     { assume asm: "N=Ax y d"
       
   692       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} = 
       
   693         Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z" 
       
   694         using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       
   695       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P})  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
       
   696         using prems
       
   697         apply -
       
   698         apply(rule a_starI)
       
   699         apply(rule al_redu)
       
   700         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   701         done
       
   702       also have "\<dots> = Cut <a>.(Cut <d>.(Cut <c>.P (y).Ax y d)  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
       
   703         using prems by simp
       
   704       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(P[c\<turnstile>c>d]) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
       
   705       proof (cases "fic P c")
       
   706         case True 
       
   707         assume "fic P c"
       
   708         then show ?thesis using prems
       
   709           apply -
       
   710           apply(rule a_starI)
       
   711           apply(rule better_CutL_intro)
       
   712           apply(rule a_Cut_l)
       
   713           apply(simp add: subst_fresh abs_fresh)
       
   714           apply(simp add: abs_fresh fresh_atm)
       
   715           apply(rule al_redu)
       
   716           apply(rule better_LAxR_intro)
       
   717           apply(simp)
       
   718           done
       
   719       next
       
   720         case False 
       
   721         assume "\<not>fic P c" 
       
   722         then show ?thesis using prems
       
   723           apply -
       
   724           apply(rule a_star_CutL)
       
   725           apply(rule a_star_CutL)
       
   726           apply(rule a_star_trans)
       
   727           apply(rule a_starI)
       
   728           apply(rule ac_redu)
       
   729           apply(rule better_left)
       
   730           apply(simp)
       
   731           apply(simp add: subst_with_ax2)
       
   732           done
       
   733       qed
       
   734       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}" using prems
       
   735         apply -
       
   736         apply(auto simp add: subst_fresh abs_fresh)
       
   737         apply(simp add: trm.inject)
       
   738         apply(simp add: alpha fresh_atm)
       
   739         apply(simp add: trm.inject)
       
   740         apply(simp add: alpha)
       
   741         apply(rule sym)
       
   742         apply(rule crename_swap)
       
   743         apply(simp)
       
   744         done
       
   745       finally 
       
   746       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a* 
       
   747                (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}"
       
   748         by simp
       
   749     }
       
   750     ultimately show ?thesis by blast
       
   751   qed
       
   752 qed
       
   753 
       
   754 lemma l_redu_subst2:
       
   755   assumes a: "M \<longrightarrow>\<^isub>l M'" 
       
   756   shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
       
   757 using a
       
   758 proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct)
       
   759   case LAxR
       
   760   then show ?case
       
   761     apply -
       
   762     apply(rule aux3)
       
   763     apply(rule better_Cut_substc)
       
   764     apply(simp add: abs_fresh)
       
   765     apply(simp add: abs_fresh)
       
   766     apply(simp add: trm.inject fresh_atm)
       
   767     apply(auto)
       
   768     apply(rule aux4)
       
   769     apply(rule sym)
       
   770     apply(rule fic_substc_crename)
       
   771     apply(simp_all)
       
   772     apply(rule a_starI)
       
   773     apply(rule al_redu)
       
   774     apply(rule aux2)
       
   775     apply(rule l_redu.intros)
       
   776     apply(simp add: subst_fresh)
       
   777     apply(simp add: fresh_atm)
       
   778     apply(rule fic_subst1)
       
   779     apply(simp_all)
       
   780     apply(rule subst_comm')
       
   781     apply(simp_all)
       
   782     done
       
   783 next
       
   784   case LAxL
       
   785   then show ?case
       
   786     apply -
       
   787     apply(rule aux3)
       
   788     apply(rule better_Cut_substc)
       
   789     apply(simp)
       
   790     apply(simp add: abs_fresh)
       
   791     apply(simp add: fresh_atm)
       
   792     apply(auto)
       
   793     apply(rule aux4)
       
   794     apply(simp add: trm.inject alpha calc_atm fresh_atm)
       
   795     apply(rule a_star_trans)
       
   796     apply(rule a_starI)
       
   797     apply(rule al_redu)
       
   798     apply(rule l_redu.intros)
       
   799     apply(simp add: subst_fresh)
       
   800     apply(simp add: fresh_atm)
       
   801     apply(rule fin_subst2)
       
   802     apply(simp_all)
       
   803     apply(rule aux4)
       
   804     apply(rule subst_comm')
       
   805     apply(simp_all)
       
   806     done
       
   807 next
       
   808   case (LNot v M N u a b)
       
   809   then show ?case
       
   810   proof -
       
   811     { assume asm: "M\<noteq>Ax u c"
       
   812       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} = 
       
   813         (Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using prems
       
   814         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   815       also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems
       
   816         by (auto intro: l_redu.intros simp add: subst_fresh)
       
   817       also have "\<dots> = (Cut <b>.N (u).M){c:=(y).P}" using prems 
       
   818         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   819       finally have ?thesis by auto
       
   820     }
       
   821     moreover
       
   822     { assume asm: "M=Ax u c"
       
   823       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} = 
       
   824         (Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using prems
       
   825         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   826       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems
       
   827         apply -
       
   828         apply(rule a_starI)
       
   829         apply(rule al_redu)
       
   830         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   831         done
       
   832       also have "\<dots> = (Cut <b>.(N{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P))" using prems
       
   833         by simp
       
   834       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P})  (u).(P[y\<turnstile>n>u]))" 
       
   835       proof (cases "fin P y")
       
   836         case True 
       
   837         assume "fin P y"
       
   838         then show ?thesis using prems
       
   839           apply -
       
   840           apply(rule a_starI)
       
   841           apply(rule better_CutR_intro)
       
   842           apply(rule al_redu)
       
   843           apply(rule better_LAxL_intro)
       
   844           apply(simp)
       
   845           done
       
   846       next
       
   847         case False 
       
   848         assume "\<not>fin P y" 
       
   849         then show ?thesis
       
   850           apply -
       
   851           apply(rule a_star_CutR)
       
   852           apply(rule a_star_trans)
       
   853           apply(rule a_starI)
       
   854           apply(rule ac_redu)
       
   855           apply(rule better_right)
       
   856           apply(simp)
       
   857           apply(simp add: subst_with_ax1)
       
   858           done
       
   859       qed
       
   860       also have "\<dots> = (Cut <b>.N (u).M){c:=(y).P}" using prems
       
   861         apply -
       
   862         apply(auto simp add: subst_fresh abs_fresh)
       
   863         apply(simp add: trm.inject)
       
   864         apply(simp add: alpha fresh_atm)
       
   865         apply(rule sym)
       
   866         apply(rule nrename_swap)
       
   867         apply(simp)
       
   868         done
       
   869       finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){c:=(y).P}" 
       
   870         by simp
       
   871     }
       
   872     ultimately show ?thesis by blast
       
   873   qed
       
   874 next
       
   875   case (LAnd1 b a1 M1 a2 M2 N z u)
       
   876   then show ?case
       
   877   proof -
       
   878     { assume asm: "N\<noteq>Ax u c"
       
   879       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} = 
       
   880         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" 
       
   881         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   882       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
       
   883         using prems
       
   884         apply -
       
   885         apply(rule a_starI)
       
   886         apply(rule al_redu)
       
   887         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   888         done
       
   889       also have "\<dots> = (Cut <a1>.M1 (u).N){c:=(y).P}" using prems 
       
   890         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   891       finally 
       
   892       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
       
   893         by simp
       
   894     } 
       
   895     moreover
       
   896     { assume asm: "N=Ax u c"
       
   897       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} = 
       
   898         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" 
       
   899         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   900       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
       
   901         using prems
       
   902         apply -
       
   903         apply(rule a_starI)
       
   904         apply(rule al_redu)
       
   905         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   906         done
       
   907       also have "\<dots> = Cut <a1>.(M1{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)" 
       
   908         using prems by simp
       
   909       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
       
   910       proof (cases "fin P y")
       
   911         case True 
       
   912         assume "fin P y"
       
   913         then show ?thesis using prems
       
   914           apply -
       
   915           apply(rule a_starI)
       
   916           apply(rule better_CutR_intro)
       
   917           apply(rule al_redu)
       
   918           apply(rule better_LAxL_intro)
       
   919           apply(simp)
       
   920           done
       
   921       next
       
   922         case False 
       
   923         assume "\<not>fin P y" 
       
   924         then show ?thesis
       
   925           apply -
       
   926           apply(rule a_star_CutR)
       
   927           apply(rule a_star_trans)
       
   928           apply(rule a_starI)
       
   929           apply(rule ac_redu)
       
   930           apply(rule better_right)
       
   931           apply(simp)
       
   932           apply(simp add: subst_with_ax1)
       
   933           done
       
   934       qed
       
   935       also have "\<dots> = (Cut <a1>.M1 (u).N){c:=(y).P}" using prems
       
   936         apply -
       
   937         apply(auto simp add: subst_fresh abs_fresh)
       
   938         apply(simp add: trm.inject)
       
   939         apply(simp add: alpha fresh_atm)
       
   940         apply(rule sym)
       
   941         apply(rule nrename_swap)
       
   942         apply(simp)
       
   943         done
       
   944       finally 
       
   945       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
       
   946         by simp
       
   947     }
       
   948     ultimately show ?thesis by blast
       
   949   qed
       
   950 next
       
   951   case (LAnd2 b a1 M1 a2 M2 N z u)
       
   952   then show ?case
       
   953   proof -
       
   954     { assume asm: "N\<noteq>Ax u c"
       
   955       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} = 
       
   956         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" 
       
   957         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   958       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
       
   959         using prems
       
   960         apply -
       
   961         apply(rule a_starI)
       
   962         apply(rule al_redu)
       
   963         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   964         done
       
   965       also have "\<dots> = (Cut <a2>.M2 (u).N){c:=(y).P}" using prems 
       
   966         by (simp add: subst_fresh abs_fresh fresh_atm)
       
   967       finally 
       
   968       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
       
   969         by simp
       
   970     } 
       
   971     moreover
       
   972     { assume asm: "N=Ax u c"
       
   973       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} = 
       
   974         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" 
       
   975         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
   976       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
       
   977         using prems
       
   978         apply -
       
   979         apply(rule a_starI)
       
   980         apply(rule al_redu)
       
   981         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
   982         done
       
   983       also have "\<dots> = Cut <a2>.(M2{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)" 
       
   984         using prems by simp
       
   985       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
       
   986       proof (cases "fin P y")
       
   987         case True 
       
   988         assume "fin P y"
       
   989         then show ?thesis using prems
       
   990           apply -
       
   991           apply(rule a_starI)
       
   992           apply(rule better_CutR_intro)
       
   993           apply(rule al_redu)
       
   994           apply(rule better_LAxL_intro)
       
   995           apply(simp)
       
   996           done
       
   997       next
       
   998         case False 
       
   999         assume "\<not>fin P y" 
       
  1000         then show ?thesis
       
  1001           apply -
       
  1002           apply(rule a_star_CutR)
       
  1003           apply(rule a_star_trans)
       
  1004           apply(rule a_starI)
       
  1005           apply(rule ac_redu)
       
  1006           apply(rule better_right)
       
  1007           apply(simp)
       
  1008           apply(simp add: subst_with_ax1)
       
  1009           done
       
  1010       qed
       
  1011       also have "\<dots> = (Cut <a2>.M2 (u).N){c:=(y).P}" using prems
       
  1012         apply -
       
  1013         apply(auto simp add: subst_fresh abs_fresh)
       
  1014         apply(simp add: trm.inject)
       
  1015         apply(simp add: alpha fresh_atm)
       
  1016         apply(rule sym)
       
  1017         apply(rule nrename_swap)
       
  1018         apply(simp)
       
  1019         done
       
  1020       finally 
       
  1021       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
       
  1022         by simp
       
  1023     }
       
  1024     ultimately show ?thesis by blast
       
  1025   qed
       
  1026 next
       
  1027   case (LOr1 b a M N1 N2 z x1 x2 y c P)
       
  1028   then show ?case
       
  1029   proof -
       
  1030     { assume asm: "N1\<noteq>Ax x1 c"
       
  1031       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
       
  1032         Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
       
  1033         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
  1034       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
       
  1035         using prems
       
  1036         apply -
       
  1037         apply(rule a_starI)
       
  1038         apply(rule al_redu)
       
  1039         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1040         done
       
  1041       also have "\<dots> = (Cut <a>.M (x1).N1){c:=(y).P}" using prems 
       
  1042         by (simp add: subst_fresh abs_fresh fresh_atm)
       
  1043       finally 
       
  1044       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
       
  1045         by simp
       
  1046     } 
       
  1047     moreover
       
  1048     { assume asm: "N1=Ax x1 c"
       
  1049       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
       
  1050         Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
       
  1051         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
  1052       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
       
  1053         using prems
       
  1054         apply -
       
  1055         apply(rule a_starI)
       
  1056         apply(rule al_redu)
       
  1057         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1058         done
       
  1059       also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x1).(Cut <c>.(Ax x1 c) (y).P)" 
       
  1060         using prems by simp
       
  1061       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P})   (x1).(P[y\<turnstile>n>x1])"
       
  1062       proof (cases "fin P y")
       
  1063         case True 
       
  1064         assume "fin P y"
       
  1065         then show ?thesis using prems
       
  1066           apply -
       
  1067           apply(rule a_starI)
       
  1068           apply(rule better_CutR_intro)
       
  1069           apply(rule al_redu)
       
  1070           apply(rule better_LAxL_intro)
       
  1071           apply(simp)
       
  1072           done
       
  1073       next
       
  1074         case False 
       
  1075         assume "\<not>fin P y" 
       
  1076         then show ?thesis
       
  1077           apply -
       
  1078           apply(rule a_star_CutR)
       
  1079           apply(rule a_star_trans)
       
  1080           apply(rule a_starI)
       
  1081           apply(rule ac_redu)
       
  1082           apply(rule better_right)
       
  1083           apply(simp)
       
  1084           apply(simp add: subst_with_ax1)
       
  1085           done
       
  1086       qed
       
  1087       also have "\<dots> = (Cut <a>.M (x1).N1){c:=(y).P}" using prems
       
  1088         apply -
       
  1089         apply(auto simp add: subst_fresh abs_fresh)
       
  1090         apply(simp add: trm.inject)
       
  1091         apply(simp add: alpha fresh_atm)
       
  1092         apply(rule sym)
       
  1093         apply(rule nrename_swap)
       
  1094         apply(simp)
       
  1095         done
       
  1096       finally 
       
  1097       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
       
  1098         by simp
       
  1099     }
       
  1100     ultimately show ?thesis by blast
       
  1101   qed
       
  1102 next
       
  1103   case (LOr2 b a M N1 N2 z x1 x2 y c P)
       
  1104   then show ?case
       
  1105   proof -
       
  1106     { assume asm: "N2\<noteq>Ax x2 c"
       
  1107       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
       
  1108         Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
       
  1109         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
  1110       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
       
  1111         using prems
       
  1112         apply -
       
  1113         apply(rule a_starI)
       
  1114         apply(rule al_redu)
       
  1115         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1116         done
       
  1117       also have "\<dots> = (Cut <a>.M (x2).N2){c:=(y).P}" using prems 
       
  1118         by (simp add: subst_fresh abs_fresh fresh_atm)
       
  1119       finally 
       
  1120       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
       
  1121         by simp
       
  1122     } 
       
  1123     moreover
       
  1124     { assume asm: "N2=Ax x2 c"
       
  1125       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
       
  1126         Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
       
  1127         using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       
  1128       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
       
  1129         using prems
       
  1130         apply -
       
  1131         apply(rule a_starI)
       
  1132         apply(rule al_redu)
       
  1133         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1134         done
       
  1135       also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x2).(Cut <c>.(Ax x2 c) (y).P)" 
       
  1136         using prems by simp
       
  1137       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(P[y\<turnstile>n>x2])"
       
  1138       proof (cases "fin P y")
       
  1139         case True 
       
  1140         assume "fin P y"
       
  1141         then show ?thesis using prems
       
  1142           apply -
       
  1143           apply(rule a_starI)
       
  1144           apply(rule better_CutR_intro)
       
  1145           apply(rule al_redu)
       
  1146           apply(rule better_LAxL_intro)
       
  1147           apply(simp)
       
  1148           done
       
  1149       next
       
  1150         case False 
       
  1151         assume "\<not>fin P y" 
       
  1152         then show ?thesis
       
  1153           apply -
       
  1154           apply(rule a_star_CutR)
       
  1155           apply(rule a_star_trans)
       
  1156           apply(rule a_starI)
       
  1157           apply(rule ac_redu)
       
  1158           apply(rule better_right)
       
  1159           apply(simp)
       
  1160           apply(simp add: subst_with_ax1)
       
  1161           done
       
  1162       qed
       
  1163       also have "\<dots> = (Cut <a>.M (x2).N2){c:=(y).P}" using prems
       
  1164         apply -
       
  1165         apply(auto simp add: subst_fresh abs_fresh)
       
  1166         apply(simp add: trm.inject)
       
  1167         apply(simp add: alpha fresh_atm)
       
  1168         apply(rule sym)
       
  1169         apply(rule nrename_swap)
       
  1170         apply(simp)
       
  1171         done
       
  1172       finally 
       
  1173       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
       
  1174         by simp
       
  1175     }
       
  1176     ultimately show ?thesis by blast
       
  1177   qed
       
  1178 next
       
  1179   case (LImp z N u Q x M b a d y c P)
       
  1180   then show ?case
       
  1181   proof -
       
  1182     { assume asm: "M\<noteq>Ax x c \<and> Q\<noteq>Ax u c"
       
  1183       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
       
  1184         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
       
  1185         using prems by (simp add: fresh_prod abs_fresh fresh_atm)
       
  1186       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
       
  1187         using prems
       
  1188         apply -
       
  1189         apply(rule a_starI)
       
  1190         apply(rule al_redu)
       
  1191         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1192         done
       
  1193       also have "\<dots> = (Cut <a>.(Cut <d>.N  (x).M) (u).Q){c:=(y).P}" using prems 
       
  1194         by (simp add: subst_fresh abs_fresh fresh_atm)
       
  1195       finally 
       
  1196       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
       
  1197                      (Cut <a>.(Cut <d>.N  (x).M) (u).Q){c:=(y).P}"
       
  1198         by simp
       
  1199     } 
       
  1200     moreover
       
  1201     { assume asm: "M=Ax x c \<and> Q\<noteq>Ax u c"
       
  1202       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
       
  1203         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
       
  1204         using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       
  1205       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
       
  1206         using prems
       
  1207         apply -
       
  1208         apply(rule a_starI)
       
  1209         apply(rule al_redu)
       
  1210         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1211         done
       
  1212       also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(Q{c:=(y).P})"
       
  1213         using prems by simp
       
  1214       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(P[y\<turnstile>n>x])) (u).(Q{c:=(y).P})"
       
  1215       proof (cases "fin P y")
       
  1216         case True 
       
  1217         assume "fin P y"
       
  1218         then show ?thesis using prems
       
  1219           apply -
       
  1220           apply(rule a_star_CutL)
       
  1221           apply(rule a_star_CutR)
       
  1222           apply(rule a_star_trans)
       
  1223           apply(rule a_starI)
       
  1224           apply(rule al_redu)
       
  1225           apply(rule better_LAxL_intro)
       
  1226           apply(simp)
       
  1227           apply(simp)
       
  1228           done
       
  1229       next
       
  1230         case False 
       
  1231         assume "\<not>fin P y" 
       
  1232         then show ?thesis using prems
       
  1233           apply -
       
  1234           apply(rule a_star_CutL)
       
  1235           apply(rule a_star_CutR)
       
  1236           apply(rule a_star_trans)
       
  1237           apply(rule a_starI)
       
  1238           apply(rule ac_redu)
       
  1239           apply(rule better_right)
       
  1240           apply(simp)
       
  1241           apply(simp add: subst_with_ax1)
       
  1242           done
       
  1243       qed
       
  1244       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
       
  1245         apply -
       
  1246         apply(auto simp add: subst_fresh abs_fresh)
       
  1247         apply(simp add: trm.inject)
       
  1248         apply(simp add: alpha fresh_atm)
       
  1249         apply(simp add: trm.inject)
       
  1250         apply(simp add: alpha)
       
  1251         apply(simp add: nrename_swap)
       
  1252         done
       
  1253       finally 
       
  1254       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
       
  1255                (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
       
  1256         by simp
       
  1257     }
       
  1258      moreover
       
  1259     { assume asm: "M\<noteq>Ax x c \<and> Q=Ax u c"
       
  1260       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
       
  1261         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
       
  1262         using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       
  1263       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
       
  1264         using prems
       
  1265         apply -
       
  1266         apply(rule a_starI)
       
  1267         apply(rule al_redu)
       
  1268         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1269         done
       
  1270       also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Cut <c>.Ax u c (y).P)"
       
  1271         using prems by simp
       
  1272       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(P[y\<turnstile>n>u])"
       
  1273       proof (cases "fin P y")
       
  1274         case True 
       
  1275         assume "fin P y"
       
  1276         then show ?thesis using prems
       
  1277           apply -
       
  1278           apply(rule a_star_CutR)
       
  1279           apply(rule a_starI)
       
  1280           apply(rule al_redu)
       
  1281           apply(rule better_LAxL_intro)
       
  1282           apply(simp)
       
  1283           done
       
  1284       next
       
  1285         case False 
       
  1286         assume "\<not>fin P y" 
       
  1287         then show ?thesis using prems
       
  1288           apply -
       
  1289           apply(rule a_star_CutR)
       
  1290           apply(rule a_star_trans)
       
  1291           apply(rule a_starI)
       
  1292           apply(rule ac_redu)
       
  1293           apply(rule better_right)
       
  1294           apply(simp)
       
  1295           apply(simp add: subst_with_ax1)
       
  1296           done
       
  1297       qed
       
  1298       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
       
  1299         apply -
       
  1300         apply(auto simp add: subst_fresh abs_fresh)
       
  1301         apply(simp add: trm.inject)
       
  1302         apply(simp add: alpha fresh_atm)
       
  1303         apply(simp add: nrename_swap)
       
  1304         done
       
  1305       finally 
       
  1306       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
       
  1307                (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
       
  1308         by simp
       
  1309     }
       
  1310      moreover
       
  1311     { assume asm: "M=Ax x c \<and> Q=Ax u c"
       
  1312       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
       
  1313         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
       
  1314         using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       
  1315       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
       
  1316         using prems
       
  1317         apply -
       
  1318         apply(rule a_starI)
       
  1319         apply(rule al_redu)
       
  1320         apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
       
  1321         done
       
  1322       also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(Cut <c>.Ax u c (y).P)"
       
  1323         using prems by simp
       
  1324       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(P[y\<turnstile>n>u])"
       
  1325       proof (cases "fin P y")
       
  1326         case True 
       
  1327         assume "fin P y"
       
  1328         then show ?thesis using prems
       
  1329           apply -
       
  1330           apply(rule a_star_CutR)
       
  1331           apply(rule a_starI)
       
  1332           apply(rule al_redu)
       
  1333           apply(rule better_LAxL_intro)
       
  1334           apply(simp)
       
  1335           done
       
  1336       next
       
  1337         case False 
       
  1338         assume "\<not>fin P y" 
       
  1339         then show ?thesis using prems
       
  1340           apply -
       
  1341           apply(rule a_star_CutR)
       
  1342           apply(rule a_star_trans)
       
  1343           apply(rule a_starI)
       
  1344           apply(rule ac_redu)
       
  1345           apply(rule better_right)
       
  1346           apply(simp)
       
  1347           apply(simp add: subst_with_ax1)
       
  1348           done
       
  1349       qed
       
  1350       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(P[y\<turnstile>n>x])) (u).(P[y\<turnstile>n>u])"
       
  1351       proof (cases "fin P y")
       
  1352         case True 
       
  1353         assume "fin P y"
       
  1354         then show ?thesis using prems
       
  1355           apply -
       
  1356           apply(rule a_star_CutL)
       
  1357           apply(rule a_star_CutR)
       
  1358           apply(rule a_starI)
       
  1359           apply(rule al_redu)
       
  1360           apply(rule better_LAxL_intro)
       
  1361           apply(simp)
       
  1362           done
       
  1363       next
       
  1364         case False 
       
  1365         assume "\<not>fin P y" 
       
  1366         then show ?thesis using prems
       
  1367           apply -
       
  1368           apply(rule a_star_CutL)
       
  1369           apply(rule a_star_CutR)
       
  1370           apply(rule a_star_trans)
       
  1371           apply(rule a_starI)
       
  1372           apply(rule ac_redu)
       
  1373           apply(rule better_right)
       
  1374           apply(simp)
       
  1375           apply(simp add: subst_with_ax1)
       
  1376           done
       
  1377       qed
       
  1378       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
       
  1379         apply -
       
  1380         apply(auto simp add: subst_fresh abs_fresh)
       
  1381         apply(simp add: trm.inject)
       
  1382         apply(rule conjI)
       
  1383         apply(simp add: alpha fresh_atm trm.inject)
       
  1384         apply(simp add: nrename_swap)
       
  1385         apply(simp add: alpha fresh_atm trm.inject)
       
  1386         apply(simp add: nrename_swap)
       
  1387         done
       
  1388       finally 
       
  1389       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
       
  1390                (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
       
  1391         by simp
       
  1392     }
       
  1393     ultimately show ?thesis by blast
       
  1394   qed
       
  1395 qed
       
  1396 
       
  1397 lemma a_redu_subst1:
       
  1398   assumes a: "M \<longrightarrow>\<^isub>a M'"
       
  1399   shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
       
  1400 using a
       
  1401 proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct)
       
  1402   case al_redu
       
  1403   then show ?case by (simp only: l_redu_subst1)
       
  1404 next
       
  1405   case ac_redu
       
  1406   then show ?case
       
  1407     apply -
       
  1408     apply(rule a_starI)
       
  1409     apply(rule a_redu.ac_redu)
       
  1410     apply(simp only: c_redu_subst1')
       
  1411     done
       
  1412 next
       
  1413   case (a_Cut_l a N x M M' y c P)
       
  1414   then show ?case
       
  1415     apply(simp add: subst_fresh fresh_a_redu)
       
  1416     apply(rule conjI)
       
  1417     apply(rule impI)+
       
  1418     apply(simp)
       
  1419     apply(drule ax_do_not_a_reduce)
       
  1420     apply(simp)
       
  1421     apply(rule impI)
       
  1422     apply(rule conjI)
       
  1423     apply(rule impI)
       
  1424     apply(simp)
       
  1425     apply(drule_tac x="y" in meta_spec)
       
  1426     apply(drule_tac x="c" in meta_spec)
       
  1427     apply(drule_tac x="P" in meta_spec)
       
  1428     apply(simp)
       
  1429     apply(rule a_star_trans)
       
  1430     apply(rule a_star_CutL)
       
  1431     apply(assumption)
       
  1432     apply(rule a_star_trans)
       
  1433     apply(rule_tac M'="P[c\<turnstile>c>a]" in a_star_CutL)
       
  1434     apply(case_tac "fic P c")
       
  1435     apply(rule a_starI)
       
  1436     apply(rule al_redu)
       
  1437     apply(rule better_LAxR_intro)
       
  1438     apply(simp)
       
  1439     apply(rule a_star_trans)
       
  1440     apply(rule a_starI)
       
  1441     apply(rule ac_redu)
       
  1442     apply(rule better_left)
       
  1443     apply(simp)
       
  1444     apply(rule subst_with_ax2)
       
  1445     apply(rule aux4)
       
  1446     apply(simp add: trm.inject)
       
  1447     apply(simp add: alpha fresh_atm)
       
  1448     apply(simp add: crename_swap)
       
  1449     apply(rule impI)
       
  1450     apply(rule a_star_CutL)
       
  1451     apply(auto)
       
  1452     done
       
  1453 next
       
  1454   case (a_Cut_r a N x M M' y c P)
       
  1455   then show ?case
       
  1456     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1457     apply(rule a_star_CutR)
       
  1458     apply(auto)[1]
       
  1459     apply(rule a_star_CutR)
       
  1460     apply(auto)[1]
       
  1461     done
       
  1462 next
       
  1463   case a_NotL
       
  1464   then show ?case 
       
  1465     apply(auto)
       
  1466     apply(generate_fresh "name")
       
  1467     apply(fresh_fun_simp)
       
  1468     apply(fresh_fun_simp)
       
  1469     apply(simp add: subst_fresh)
       
  1470     apply(rule a_star_CutR)
       
  1471     apply(rule a_star_NotL)
       
  1472     apply(auto)[1]
       
  1473     apply(rule a_star_NotL)
       
  1474     apply(auto)[1]
       
  1475     done
       
  1476 next
       
  1477   case a_NotR
       
  1478   then show ?case 
       
  1479     apply(auto)
       
  1480     apply(rule a_star_NotR)
       
  1481     apply(auto)[1]
       
  1482     done
       
  1483 next
       
  1484   case a_AndR_l
       
  1485   then show ?case
       
  1486     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1487     apply(rule a_star_AndR)
       
  1488     apply(auto)
       
  1489     done
       
  1490 next
       
  1491   case a_AndR_r
       
  1492   then show ?case
       
  1493     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1494     apply(rule a_star_AndR)
       
  1495     apply(auto)
       
  1496     done
       
  1497 next
       
  1498   case a_AndL1
       
  1499   then show ?case 
       
  1500     apply(auto)
       
  1501     apply(generate_fresh "name")
       
  1502     apply(fresh_fun_simp)
       
  1503     apply(fresh_fun_simp)
       
  1504     apply(simp add: subst_fresh)
       
  1505     apply(rule a_star_CutR)
       
  1506     apply(rule a_star_AndL1)
       
  1507     apply(auto)[1]
       
  1508     apply(rule a_star_AndL1)
       
  1509     apply(auto)[1]
       
  1510     done
       
  1511 next
       
  1512   case a_AndL2
       
  1513   then show ?case 
       
  1514     apply(auto)
       
  1515     apply(generate_fresh "name")
       
  1516     apply(fresh_fun_simp)
       
  1517     apply(fresh_fun_simp)
       
  1518     apply(simp add: subst_fresh)
       
  1519     apply(rule a_star_CutR)
       
  1520     apply(rule a_star_AndL2)
       
  1521     apply(auto)[1]
       
  1522     apply(rule a_star_AndL2)
       
  1523     apply(auto)[1]
       
  1524     done
       
  1525 next
       
  1526   case a_OrR1
       
  1527   then show ?case
       
  1528     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1529     apply(rule a_star_OrR1)
       
  1530     apply(auto)
       
  1531     done
       
  1532 next
       
  1533   case a_OrR2
       
  1534   then show ?case
       
  1535     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1536     apply(rule a_star_OrR2)
       
  1537     apply(auto)
       
  1538     done
       
  1539 next
       
  1540   case a_OrL_l
       
  1541   then show ?case 
       
  1542     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1543     apply(generate_fresh "name")
       
  1544     apply(fresh_fun_simp)
       
  1545     apply(fresh_fun_simp)
       
  1546     apply(simp add: subst_fresh)
       
  1547     apply(rule a_star_CutR)
       
  1548     apply(rule a_star_OrL)
       
  1549     apply(auto)
       
  1550     apply(rule a_star_OrL)
       
  1551     apply(auto)
       
  1552     done
       
  1553 next
       
  1554   case a_OrL_r
       
  1555   then show ?case 
       
  1556     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1557     apply(generate_fresh "name")
       
  1558     apply(fresh_fun_simp)
       
  1559     apply(fresh_fun_simp)
       
  1560     apply(simp add: subst_fresh)
       
  1561     apply(rule a_star_CutR)
       
  1562     apply(rule a_star_OrL)
       
  1563     apply(auto)
       
  1564     apply(rule a_star_OrL)
       
  1565     apply(auto)
       
  1566     done
       
  1567 next
       
  1568   case a_ImpR
       
  1569   then show ?case
       
  1570     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1571     apply(rule a_star_ImpR)
       
  1572     apply(auto)
       
  1573     done
       
  1574 next
       
  1575   case a_ImpL_r
       
  1576   then show ?case 
       
  1577     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1578     apply(generate_fresh "name")
       
  1579     apply(fresh_fun_simp)
       
  1580     apply(fresh_fun_simp)
       
  1581     apply(simp add: subst_fresh)
       
  1582     apply(rule a_star_CutR)
       
  1583     apply(rule a_star_ImpL)
       
  1584     apply(auto)
       
  1585     apply(rule a_star_ImpL)
       
  1586     apply(auto)
       
  1587     done
       
  1588 next
       
  1589   case a_ImpL_l
       
  1590   then show ?case 
       
  1591     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1592     apply(generate_fresh "name")
       
  1593     apply(fresh_fun_simp)
       
  1594     apply(fresh_fun_simp)
       
  1595     apply(simp add: subst_fresh)
       
  1596     apply(rule a_star_CutR)
       
  1597     apply(rule a_star_ImpL)
       
  1598     apply(auto)
       
  1599     apply(rule a_star_ImpL)
       
  1600     apply(auto)
       
  1601     done
       
  1602 qed
       
  1603 
       
  1604 lemma a_redu_subst2:
       
  1605   assumes a: "M \<longrightarrow>\<^isub>a M'"
       
  1606   shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
       
  1607 using a
       
  1608 proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct)
       
  1609   case al_redu
       
  1610   then show ?case by (simp only: l_redu_subst2)
       
  1611 next
       
  1612   case ac_redu
       
  1613   then show ?case
       
  1614     apply -
       
  1615     apply(rule a_starI)
       
  1616     apply(rule a_redu.ac_redu)
       
  1617     apply(simp only: c_redu_subst2')
       
  1618     done
       
  1619 next
       
  1620   case (a_Cut_r a N x M M' y c P)
       
  1621   then show ?case
       
  1622     apply(simp add: subst_fresh fresh_a_redu)
       
  1623     apply(rule conjI)
       
  1624     apply(rule impI)+
       
  1625     apply(simp)
       
  1626     apply(drule ax_do_not_a_reduce)
       
  1627     apply(simp)
       
  1628     apply(rule impI)
       
  1629     apply(rule conjI)
       
  1630     apply(rule impI)
       
  1631     apply(simp)
       
  1632     apply(drule_tac x="c" in meta_spec)
       
  1633     apply(drule_tac x="y" in meta_spec)
       
  1634     apply(drule_tac x="P" in meta_spec)
       
  1635     apply(simp)
       
  1636     apply(rule a_star_trans)
       
  1637     apply(rule a_star_CutR)
       
  1638     apply(assumption)
       
  1639     apply(rule a_star_trans)
       
  1640     apply(rule_tac N'="P[y\<turnstile>n>x]" in a_star_CutR)
       
  1641     apply(case_tac "fin P y")
       
  1642     apply(rule a_starI)
       
  1643     apply(rule al_redu)
       
  1644     apply(rule better_LAxL_intro)
       
  1645     apply(simp)
       
  1646     apply(rule a_star_trans)
       
  1647     apply(rule a_starI)
       
  1648     apply(rule ac_redu)
       
  1649     apply(rule better_right)
       
  1650     apply(simp)
       
  1651     apply(rule subst_with_ax1)
       
  1652     apply(rule aux4)
       
  1653     apply(simp add: trm.inject)
       
  1654     apply(simp add: alpha fresh_atm)
       
  1655     apply(simp add: nrename_swap)
       
  1656     apply(rule impI)
       
  1657     apply(rule a_star_CutR)
       
  1658     apply(auto)
       
  1659     done
       
  1660 next
       
  1661   case (a_Cut_l a N x M M' y c P)
       
  1662   then show ?case
       
  1663     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1664     apply(rule a_star_CutL)
       
  1665     apply(auto)[1]
       
  1666     apply(rule a_star_CutL)
       
  1667     apply(auto)[1]
       
  1668     done
       
  1669 next
       
  1670   case a_NotR
       
  1671   then show ?case 
       
  1672     apply(auto)
       
  1673     apply(generate_fresh "coname")
       
  1674     apply(fresh_fun_simp)
       
  1675     apply(fresh_fun_simp)
       
  1676     apply(simp add: subst_fresh)
       
  1677     apply(rule a_star_CutL)
       
  1678     apply(rule a_star_NotR)
       
  1679     apply(auto)[1]
       
  1680     apply(rule a_star_NotR)
       
  1681     apply(auto)[1]
       
  1682     done
       
  1683 next
       
  1684   case a_NotL
       
  1685   then show ?case 
       
  1686     apply(auto)
       
  1687     apply(rule a_star_NotL)
       
  1688     apply(auto)[1]
       
  1689     done
       
  1690 next
       
  1691   case a_AndR_l
       
  1692   then show ?case
       
  1693     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1694     apply(generate_fresh "coname")
       
  1695     apply(fresh_fun_simp)
       
  1696     apply(fresh_fun_simp)
       
  1697     apply(simp add: subst_fresh)
       
  1698     apply(rule a_star_CutL)
       
  1699     apply(rule a_star_AndR)
       
  1700     apply(auto)
       
  1701     apply(rule a_star_AndR)
       
  1702     apply(auto)
       
  1703     done
       
  1704 next
       
  1705   case a_AndR_r
       
  1706   then show ?case
       
  1707     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1708     apply(generate_fresh "coname")
       
  1709     apply(fresh_fun_simp)
       
  1710     apply(fresh_fun_simp)
       
  1711     apply(simp add: subst_fresh)
       
  1712     apply(rule a_star_CutL)
       
  1713     apply(rule a_star_AndR)
       
  1714     apply(auto)
       
  1715     apply(rule a_star_AndR)
       
  1716     apply(auto)
       
  1717     done
       
  1718 next
       
  1719   case a_AndL1
       
  1720   then show ?case
       
  1721     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1722     apply(rule a_star_AndL1)
       
  1723     apply(auto)
       
  1724     done
       
  1725 next
       
  1726   case a_AndL2
       
  1727   then show ?case
       
  1728     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1729     apply(rule a_star_AndL2)
       
  1730     apply(auto)
       
  1731     done
       
  1732 next
       
  1733   case a_OrR1
       
  1734   then show ?case 
       
  1735     apply(auto)
       
  1736     apply(generate_fresh "coname")
       
  1737     apply(fresh_fun_simp)
       
  1738     apply(fresh_fun_simp)
       
  1739     apply(simp add: subst_fresh)
       
  1740     apply(rule a_star_CutL)
       
  1741     apply(rule a_star_OrR1)
       
  1742     apply(auto)[1]
       
  1743     apply(rule a_star_OrR1)
       
  1744     apply(auto)[1]
       
  1745     done
       
  1746 next
       
  1747   case a_OrR2
       
  1748   then show ?case 
       
  1749     apply(auto)
       
  1750     apply(generate_fresh "coname")
       
  1751     apply(fresh_fun_simp)
       
  1752     apply(fresh_fun_simp)
       
  1753     apply(simp add: subst_fresh)
       
  1754     apply(rule a_star_CutL)
       
  1755     apply(rule a_star_OrR2)
       
  1756     apply(auto)[1]
       
  1757     apply(rule a_star_OrR2)
       
  1758     apply(auto)[1]
       
  1759     done
       
  1760 next
       
  1761   case a_OrL_l
       
  1762   then show ?case
       
  1763     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1764     apply(rule a_star_OrL)
       
  1765     apply(auto)
       
  1766     done
       
  1767 next
       
  1768   case a_OrL_r
       
  1769   then show ?case
       
  1770     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1771     apply(rule a_star_OrL)
       
  1772     apply(auto)
       
  1773     done
       
  1774 next
       
  1775   case a_ImpR
       
  1776   then show ?case 
       
  1777     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1778     apply(generate_fresh "coname")
       
  1779     apply(fresh_fun_simp)
       
  1780     apply(fresh_fun_simp)
       
  1781     apply(simp add: subst_fresh)
       
  1782     apply(rule a_star_CutL)
       
  1783     apply(rule a_star_ImpR)
       
  1784     apply(auto)
       
  1785     apply(rule a_star_ImpR)
       
  1786     apply(auto)
       
  1787     done
       
  1788 next
       
  1789   case a_ImpL_l
       
  1790   then show ?case
       
  1791     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1792     apply(rule a_star_ImpL)
       
  1793     apply(auto)
       
  1794     done
       
  1795 next
       
  1796   case a_ImpL_r
       
  1797   then show ?case
       
  1798     apply(auto simp add: subst_fresh fresh_a_redu)
       
  1799     apply(rule a_star_ImpL)
       
  1800     apply(auto)
       
  1801     done
       
  1802 qed
       
  1803 
       
  1804 lemma a_star_subst1:
       
  1805   assumes a: "M \<longrightarrow>\<^isub>a* M'"
       
  1806   shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
       
  1807 using a
       
  1808 apply(induct)
       
  1809 apply(blast)
       
  1810 apply(drule_tac y="y" and c="c" and P="P" in a_redu_subst1)
       
  1811 apply(auto)
       
  1812 done
       
  1813 
       
  1814 lemma a_star_subst2:
       
  1815   assumes a: "M \<longrightarrow>\<^isub>a* M'"
       
  1816   shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
       
  1817 using a
       
  1818 apply(induct)
       
  1819 apply(blast)
       
  1820 apply(drule_tac y="y" and c="c" and P="P" in a_redu_subst2)
       
  1821 apply(auto)
       
  1822 done
       
  1823 
       
  1824 text {* Candidates and SN *}
       
  1825 
       
  1826 text {* SNa *}
       
  1827 
       
  1828 inductive 
       
  1829   SNa :: "trm \<Rightarrow> bool"
       
  1830 where
       
  1831   SNaI: "(\<And>M'. M \<longrightarrow>\<^isub>a M' \<Longrightarrow> SNa M') \<Longrightarrow> SNa M"
       
  1832 
       
  1833 lemma SNa_induct[consumes 1]:
       
  1834   assumes major: "SNa M"
       
  1835   assumes hyp: "\<And>M'. SNa M' \<Longrightarrow> (\<forall>M''. M'\<longrightarrow>\<^isub>a M'' \<longrightarrow> P M'' \<Longrightarrow> P M')"
       
  1836   shows "P M"
       
  1837   apply (rule major[THEN SNa.induct])
       
  1838   apply (rule hyp)
       
  1839   apply (rule SNaI)
       
  1840   apply (blast)+
       
  1841   done
       
  1842 
       
  1843 
       
  1844 lemma double_SNa_aux:
       
  1845   assumes a_SNa: "SNa a"
       
  1846   and b_SNa: "SNa b"
       
  1847   and hyp: "\<And>x z.
       
  1848     (\<And>y. x\<longrightarrow>\<^isub>a y \<Longrightarrow> SNa y) \<Longrightarrow>
       
  1849     (\<And>y. x\<longrightarrow>\<^isub>a y \<Longrightarrow> P y z) \<Longrightarrow>
       
  1850     (\<And>u. z\<longrightarrow>\<^isub>a u \<Longrightarrow> SNa u) \<Longrightarrow>
       
  1851     (\<And>u. z\<longrightarrow>\<^isub>a u \<Longrightarrow> P x u) \<Longrightarrow> P x z"
       
  1852   shows "P a b"
       
  1853 proof -
       
  1854   from a_SNa
       
  1855   have r: "\<And>b. SNa b \<Longrightarrow> P a b"
       
  1856   proof (induct a rule: SNa.induct)
       
  1857     case (SNaI x)
       
  1858     note SNa' = this
       
  1859     have "SNa b" by fact
       
  1860     thus ?case
       
  1861     proof (induct b rule: SNa.induct)
       
  1862       case (SNaI y)
       
  1863       show ?case
       
  1864         apply (rule hyp)
       
  1865         apply (erule SNa')
       
  1866         apply (erule SNa')
       
  1867         apply (rule SNa.SNaI)
       
  1868         apply (erule SNaI)+
       
  1869         done
       
  1870     qed
       
  1871   qed
       
  1872   from b_SNa show ?thesis by (rule r)
       
  1873 qed
       
  1874 
       
  1875 lemma double_SNa:
       
  1876   "\<lbrakk>SNa a; SNa b; \<forall>x z. ((\<forall>y. x\<longrightarrow>\<^isub>ay \<longrightarrow> P y z) \<and> (\<forall>u. z\<longrightarrow>\<^isub>a u \<longrightarrow> P x u)) \<longrightarrow> P x z\<rbrakk> \<Longrightarrow> P a b"
       
  1877 apply(rule_tac double_SNa_aux)
       
  1878 apply(assumption)+
       
  1879 apply(blast)
       
  1880 done
       
  1881 
       
  1882 lemma a_preserves_SNa:
       
  1883   assumes a: "SNa M" "M\<longrightarrow>\<^isub>a M'"
       
  1884   shows "SNa M'"
       
  1885 using a 
       
  1886 by (erule_tac SNa.cases) (simp)
       
  1887 
       
  1888 lemma a_star_preserves_SNa:
       
  1889   assumes a: "SNa M" and b: "M\<longrightarrow>\<^isub>a* M'"
       
  1890   shows "SNa M'"
       
  1891 using b a
       
  1892 by (induct) (auto simp add: a_preserves_SNa)
       
  1893 
       
  1894 lemma Ax_in_SNa:
       
  1895   shows "SNa (Ax x a)"
       
  1896 apply(rule SNaI)
       
  1897 apply(erule a_redu.cases, auto)
       
  1898 apply(erule l_redu.cases, auto)
       
  1899 apply(erule c_redu.cases, auto)
       
  1900 done
       
  1901 
       
  1902 lemma NotL_in_SNa:
       
  1903   assumes a: "SNa M"
       
  1904   shows "SNa (NotL <a>.M x)"
       
  1905 using a
       
  1906 apply(induct)
       
  1907 apply(rule SNaI)
       
  1908 apply(erule a_redu.cases, auto)
       
  1909 apply(erule l_redu.cases, auto)
       
  1910 apply(erule c_redu.cases, auto)
       
  1911 apply(auto simp add: trm.inject alpha)
       
  1912 apply(rotate_tac 1)
       
  1913 apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
       
  1914 apply(simp add: a_redu.eqvt)
       
  1915 apply(subgoal_tac "NotL <a>.([(a,aa)]\<bullet>M'a) x = NotL <aa>.M'a x")
       
  1916 apply(simp)
       
  1917 apply(simp add: trm.inject alpha fresh_a_redu)
       
  1918 done
       
  1919 
       
  1920 lemma NotR_in_SNa:
       
  1921   assumes a: "SNa M"
       
  1922   shows "SNa (NotR (x).M a)"
       
  1923 using a
       
  1924 apply(induct)
       
  1925 apply(rule SNaI)
       
  1926 apply(erule a_redu.cases, auto)
       
  1927 apply(erule l_redu.cases, auto)
       
  1928 apply(erule c_redu.cases, auto)
       
  1929 apply(auto simp add: trm.inject alpha)
       
  1930 apply(rotate_tac 1)
       
  1931 apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
       
  1932 apply(simp add: a_redu.eqvt)
       
  1933 apply(rule_tac s="NotR (x).([(x,xa)]\<bullet>M'a) a" in subst)
       
  1934 apply(simp add: trm.inject alpha fresh_a_redu)
       
  1935 apply(simp)
       
  1936 done
       
  1937 
       
  1938 lemma AndL1_in_SNa:
       
  1939   assumes a: "SNa M"
       
  1940   shows "SNa (AndL1 (x).M y)"
       
  1941 using a
       
  1942 apply(induct)
       
  1943 apply(rule SNaI)
       
  1944 apply(erule a_redu.cases, auto)
       
  1945 apply(erule l_redu.cases, auto)
       
  1946 apply(erule c_redu.cases, auto)
       
  1947 apply(auto simp add: trm.inject alpha)
       
  1948 apply(rotate_tac 1)
       
  1949 apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
       
  1950 apply(simp add: a_redu.eqvt)
       
  1951 apply(rule_tac s="AndL1 x.([(x,xa)]\<bullet>M'a) y" in subst)
       
  1952 apply(simp add: trm.inject alpha fresh_a_redu)
       
  1953 apply(simp)
       
  1954 done
       
  1955 
       
  1956 lemma AndL2_in_SNa:
       
  1957   assumes a: "SNa M"
       
  1958   shows "SNa (AndL2 (x).M y)"
       
  1959 using a
       
  1960 apply(induct)
       
  1961 apply(rule SNaI)
       
  1962 apply(erule a_redu.cases, auto)
       
  1963 apply(erule l_redu.cases, auto)
       
  1964 apply(erule c_redu.cases, auto)
       
  1965 apply(auto simp add: trm.inject alpha)
       
  1966 apply(rotate_tac 1)
       
  1967 apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
       
  1968 apply(simp add: a_redu.eqvt)
       
  1969 apply(rule_tac s="AndL2 x.([(x,xa)]\<bullet>M'a) y" in subst)
       
  1970 apply(simp add: trm.inject alpha fresh_a_redu)
       
  1971 apply(simp)
       
  1972 done
       
  1973 
       
  1974 lemma OrR1_in_SNa:
       
  1975   assumes a: "SNa M"
       
  1976   shows "SNa (OrR1 <a>.M b)"
       
  1977 using a
       
  1978 apply(induct)
       
  1979 apply(rule SNaI)
       
  1980 apply(erule a_redu.cases, auto)
       
  1981 apply(erule l_redu.cases, auto)
       
  1982 apply(erule c_redu.cases, auto)
       
  1983 apply(auto simp add: trm.inject alpha)
       
  1984 apply(rotate_tac 1)
       
  1985 apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
       
  1986 apply(simp add: a_redu.eqvt)
       
  1987 apply(rule_tac s="OrR1 <a>.([(a,aa)]\<bullet>M'a) b" in subst)
       
  1988 apply(simp add: trm.inject alpha fresh_a_redu)
       
  1989 apply(simp)
       
  1990 done
       
  1991 
       
  1992 lemma OrR2_in_SNa:
       
  1993   assumes a: "SNa M"
       
  1994   shows "SNa (OrR2 <a>.M b)"
       
  1995 using a
       
  1996 apply(induct)
       
  1997 apply(rule SNaI)
       
  1998 apply(erule a_redu.cases, auto)
       
  1999 apply(erule l_redu.cases, auto)
       
  2000 apply(erule c_redu.cases, auto)
       
  2001 apply(auto simp add: trm.inject alpha)
       
  2002 apply(rotate_tac 1)
       
  2003 apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
       
  2004 apply(simp add: a_redu.eqvt)
       
  2005 apply(rule_tac s="OrR2 <a>.([(a,aa)]\<bullet>M'a) b" in subst)
       
  2006 apply(simp add: trm.inject alpha fresh_a_redu)
       
  2007 apply(simp)
       
  2008 done
       
  2009 
       
  2010 lemma ImpR_in_SNa:
       
  2011   assumes a: "SNa M"
       
  2012   shows "SNa (ImpR (x).<a>.M b)"
       
  2013 using a
       
  2014 apply(induct)
       
  2015 apply(rule SNaI)
       
  2016 apply(erule a_redu.cases, auto)
       
  2017 apply(erule l_redu.cases, auto)
       
  2018 apply(erule c_redu.cases, auto)
       
  2019 apply(auto simp add: trm.inject alpha abs_fresh abs_perm calc_atm)
       
  2020 apply(rotate_tac 1)
       
  2021 apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
       
  2022 apply(simp add: a_redu.eqvt)
       
  2023 apply(rule_tac s="ImpR (x).<a>.([(a,aa)]\<bullet>M'a) b" in subst)
       
  2024 apply(simp add: trm.inject alpha fresh_a_redu)
       
  2025 apply(simp)
       
  2026 apply(rotate_tac 1)
       
  2027 apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
       
  2028 apply(simp add: a_redu.eqvt)
       
  2029 apply(rule_tac s="ImpR (x).<a>.([(x,xa)]\<bullet>M'a) b" in subst)
       
  2030 apply(simp add: trm.inject alpha fresh_a_redu abs_fresh abs_perm calc_atm)
       
  2031 apply(simp)
       
  2032 apply(rotate_tac 1)
       
  2033 apply(drule_tac x="[(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a" in meta_spec)
       
  2034 apply(simp add: a_redu.eqvt)
       
  2035 apply(rule_tac s="ImpR (x).<a>.([(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a) b" in subst)
       
  2036 apply(simp add: trm.inject alpha fresh_a_redu abs_fresh abs_perm calc_atm)
       
  2037 apply(simp add: fresh_left calc_atm fresh_a_redu)
       
  2038 apply(simp)
       
  2039 done
       
  2040 
       
  2041 lemma AndR_in_SNa:
       
  2042   assumes a: "SNa M" "SNa N"
       
  2043   shows "SNa (AndR <a>.M <b>.N c)"
       
  2044 apply(rule_tac a="M" and b="N" in double_SNa)
       
  2045 apply(rule a)+
       
  2046 apply(auto)
       
  2047 apply(rule SNaI)
       
  2048 apply(drule a_redu_AndR_elim)
       
  2049 apply(auto)
       
  2050 done
       
  2051 
       
  2052 lemma OrL_in_SNa:
       
  2053   assumes a: "SNa M" "SNa N"
       
  2054   shows "SNa (OrL (x).M (y).N z)"
       
  2055 apply(rule_tac a="M" and b="N" in double_SNa)
       
  2056 apply(rule a)+
       
  2057 apply(auto)
       
  2058 apply(rule SNaI)
       
  2059 apply(drule a_redu_OrL_elim)
       
  2060 apply(auto)
       
  2061 done
       
  2062 
       
  2063 lemma ImpL_in_SNa:
       
  2064   assumes a: "SNa M" "SNa N"
       
  2065   shows "SNa (ImpL <a>.M (y).N z)"
       
  2066 apply(rule_tac a="M" and b="N" in double_SNa)
       
  2067 apply(rule a)+
       
  2068 apply(auto)
       
  2069 apply(rule SNaI)
       
  2070 apply(drule a_redu_ImpL_elim)
       
  2071 apply(auto)
       
  2072 done
       
  2073 
       
  2074 lemma SNa_eqvt:
       
  2075   fixes pi1::"name prm"
       
  2076   and   pi2::"coname prm"
       
  2077   shows "SNa M \<Longrightarrow> SNa (pi1\<bullet>M)"
       
  2078   and   "SNa M \<Longrightarrow> SNa (pi2\<bullet>M)"
       
  2079 apply -
       
  2080 apply(induct rule: SNa.induct)
       
  2081 apply(rule SNaI)
       
  2082 apply(drule_tac pi="(rev pi1)" in a_redu.eqvt(1))
       
  2083 apply(rotate_tac 1)
       
  2084 apply(drule_tac x="(rev pi1)\<bullet>M'" in meta_spec)
       
  2085 apply(perm_simp)
       
  2086 apply(induct rule: SNa.induct)
       
  2087 apply(rule SNaI)
       
  2088 apply(drule_tac pi="(rev pi2)" in a_redu.eqvt(2))
       
  2089 apply(rotate_tac 1)
       
  2090 apply(drule_tac x="(rev pi2)\<bullet>M'" in meta_spec)
       
  2091 apply(perm_simp)
       
  2092 done
       
  2093 
       
  2094 text {* set operators *}
       
  2095 
       
  2096 definition AXIOMSn :: "ty \<Rightarrow> ntrm set" where
       
  2097   "AXIOMSn B \<equiv> { (x):(Ax y b) | x y b. True }"
       
  2098 
       
  2099 definition AXIOMSc::"ty \<Rightarrow> ctrm set" where
       
  2100   "AXIOMSc B \<equiv> { <a>:(Ax y b) | a y b. True }"
       
  2101 
       
  2102 definition BINDINGn::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set" where
       
  2103   "BINDINGn B X \<equiv> { (x):M | x M. \<forall>a P. <a>:P\<in>X \<longrightarrow> SNa (M{x:=<a>.P})}"
       
  2104 
       
  2105 definition BINDINGc::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set" where
       
  2106   "BINDINGc B X \<equiv> { <a>:M | a M. \<forall>x P. (x):P\<in>X \<longrightarrow> SNa (M{a:=(x).P})}"
       
  2107 
       
  2108 lemma BINDINGn_decreasing:
       
  2109   shows "X\<subseteq>Y \<Longrightarrow> BINDINGn B Y \<subseteq> BINDINGn B X"
       
  2110 by (simp add: BINDINGn_def) (blast) 
       
  2111 
       
  2112 lemma BINDINGc_decreasing:
       
  2113   shows "X\<subseteq>Y \<Longrightarrow> BINDINGc B Y \<subseteq> BINDINGc B X"
       
  2114 by (simp add: BINDINGc_def) (blast) 
       
  2115   
       
  2116 nominal_primrec
       
  2117   NOTRIGHT :: "ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set"
       
  2118 where
       
  2119  "NOTRIGHT (NOT B) X = { <a>:NotR (x).M a | a x M. fic (NotR (x).M a) a \<and> (x):M \<in> X }"
       
  2120 apply(rule TrueI)+
       
  2121 done
       
  2122 
       
  2123 lemma NOTRIGHT_eqvt_name:
       
  2124   fixes pi::"name prm"
       
  2125   shows "(pi\<bullet>(NOTRIGHT (NOT B) X)) = NOTRIGHT (NOT B) (pi\<bullet>X)"
       
  2126 apply(auto simp add: perm_set_eq)
       
  2127 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2128 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2129 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2130 apply(simp)
       
  2131 apply(rule conjI)
       
  2132 apply(drule_tac pi="pi" in fic.eqvt(1))
       
  2133 apply(simp)
       
  2134 apply(rule_tac x="(xb):M" in exI)
       
  2135 apply(simp)
       
  2136 apply(rule_tac x="(rev pi)\<bullet>(<a>:NotR (xa).M a)" in exI)
       
  2137 apply(perm_simp)
       
  2138 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2139 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2140 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2141 apply(simp add: swap_simps)
       
  2142 apply(drule_tac pi="rev pi" in fic.eqvt(1))
       
  2143 apply(simp)
       
  2144 apply(drule sym)
       
  2145 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2146 apply(simp add: swap_simps)
       
  2147 done
       
  2148 
       
  2149 lemma NOTRIGHT_eqvt_coname:
       
  2150   fixes pi::"coname prm"
       
  2151   shows "(pi\<bullet>(NOTRIGHT (NOT B) X)) = NOTRIGHT (NOT B) (pi\<bullet>X)"
       
  2152 apply(auto simp add: perm_set_eq)
       
  2153 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2154 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2155 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2156 apply(simp)
       
  2157 apply(rule conjI)
       
  2158 apply(drule_tac pi="pi" in fic.eqvt(2))
       
  2159 apply(simp)
       
  2160 apply(rule_tac x="(xb):M" in exI)
       
  2161 apply(simp)
       
  2162 apply(rule_tac x="<((rev pi)\<bullet>a)>:NotR ((rev pi)\<bullet>xa).((rev pi)\<bullet>M) ((rev pi)\<bullet>a)" in exI)
       
  2163 apply(perm_simp)
       
  2164 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2165 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2166 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2167 apply(simp add: swap_simps)
       
  2168 apply(drule_tac pi="rev pi" in fic.eqvt(2))
       
  2169 apply(simp)
       
  2170 apply(drule sym)
       
  2171 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2172 apply(simp add: swap_simps)
       
  2173 done
       
  2174   
       
  2175 nominal_primrec
       
  2176   NOTLEFT :: "ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set"
       
  2177 where
       
  2178  "NOTLEFT (NOT B) X = { (x):NotL <a>.M x | a x M. fin (NotL <a>.M x) x \<and> <a>:M \<in> X }"
       
  2179 apply(rule TrueI)+
       
  2180 done
       
  2181 
       
  2182 lemma NOTLEFT_eqvt_name:
       
  2183   fixes pi::"name prm"
       
  2184   shows "(pi\<bullet>(NOTLEFT (NOT B) X)) = NOTLEFT (NOT B) (pi\<bullet>X)"
       
  2185 apply(auto simp add: perm_set_eq)
       
  2186 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2187 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2188 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2189 apply(simp)
       
  2190 apply(rule conjI)
       
  2191 apply(drule_tac pi="pi" in fin.eqvt(1))
       
  2192 apply(simp)
       
  2193 apply(rule_tac x="<a>:M" in exI)
       
  2194 apply(simp)
       
  2195 apply(rule_tac x="(((rev pi)\<bullet>xa)):NotL <((rev pi)\<bullet>a)>.((rev pi)\<bullet>M) ((rev pi)\<bullet>xa)" in exI)
       
  2196 apply(perm_simp)
       
  2197 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2198 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2199 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2200 apply(simp add: swap_simps)
       
  2201 apply(drule_tac pi="rev pi" in fin.eqvt(1))
       
  2202 apply(simp)
       
  2203 apply(drule sym)
       
  2204 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2205 apply(simp add: swap_simps)
       
  2206 done
       
  2207 
       
  2208 lemma NOTLEFT_eqvt_coname:
       
  2209   fixes pi::"coname prm"
       
  2210   shows "(pi\<bullet>(NOTLEFT (NOT B) X)) = NOTLEFT (NOT B) (pi\<bullet>X)"
       
  2211 apply(auto simp add: perm_set_eq)
       
  2212 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2213 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2214 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2215 apply(simp)
       
  2216 apply(rule conjI)
       
  2217 apply(drule_tac pi="pi" in fin.eqvt(2))
       
  2218 apply(simp)
       
  2219 apply(rule_tac x="<a>:M" in exI)
       
  2220 apply(simp)
       
  2221 apply(rule_tac x="(((rev pi)\<bullet>xa)):NotL <((rev pi)\<bullet>a)>.((rev pi)\<bullet>M) ((rev pi)\<bullet>xa)" in exI)
       
  2222 apply(perm_simp)
       
  2223 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2224 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2225 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2226 apply(simp add: swap_simps)
       
  2227 apply(drule_tac pi="rev pi" in fin.eqvt(2))
       
  2228 apply(simp)
       
  2229 apply(drule sym)
       
  2230 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2231 apply(simp add: swap_simps)
       
  2232 done
       
  2233   
       
  2234 nominal_primrec
       
  2235   ANDRIGHT :: "ty \<Rightarrow> ctrm set \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
       
  2236 where
       
  2237  "ANDRIGHT (B AND C) X Y = 
       
  2238             { <c>:AndR <a>.M <b>.N c | c a b M N. fic (AndR <a>.M <b>.N c) c \<and> <a>:M \<in> X \<and> <b>:N \<in> Y }"
       
  2239 apply(rule TrueI)+
       
  2240 done
       
  2241 
       
  2242 lemma ANDRIGHT_eqvt_name:
       
  2243   fixes pi::"name prm"
       
  2244   shows "(pi\<bullet>(ANDRIGHT (A AND B) X Y)) = ANDRIGHT (A AND B) (pi\<bullet>X) (pi\<bullet>Y)"
       
  2245 apply(auto simp add: perm_set_eq)
       
  2246 apply(rule_tac x="pi\<bullet>c" in exI)
       
  2247 apply(rule_tac x="pi\<bullet>a" in exI)
       
  2248 apply(rule_tac x="pi\<bullet>b" in exI)
       
  2249 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2250 apply(rule_tac x="pi\<bullet>N" in exI)
       
  2251 apply(simp)
       
  2252 apply(rule conjI)
       
  2253 apply(drule_tac pi="pi" in fic.eqvt(1))
       
  2254 apply(simp)
       
  2255 apply(rule conjI)
       
  2256 apply(rule_tac x="<a>:M" in exI)
       
  2257 apply(simp)
       
  2258 apply(rule_tac x="<b>:N" in exI)
       
  2259 apply(simp)
       
  2260 apply(rule_tac x="(rev pi)\<bullet>(<c>:AndR <a>.M <b>.N c)" in exI)
       
  2261 apply(perm_simp)
       
  2262 apply(rule_tac x="(rev pi)\<bullet>c" in exI)
       
  2263 apply(rule_tac x="(rev pi)\<bullet>a" in exI)
       
  2264 apply(rule_tac x="(rev pi)\<bullet>b" in exI)
       
  2265 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2266 apply(rule_tac x="(rev pi)\<bullet>N" in exI)
       
  2267 apply(simp add: swap_simps)
       
  2268 apply(drule_tac pi="rev pi" in fic.eqvt(1))
       
  2269 apply(simp)
       
  2270 apply(drule sym)
       
  2271 apply(drule sym)
       
  2272 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2273 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2274 apply(simp add: swap_simps)
       
  2275 done
       
  2276 
       
  2277 lemma ANDRIGHT_eqvt_coname:
       
  2278   fixes pi::"coname prm"
       
  2279   shows "(pi\<bullet>(ANDRIGHT (A AND B) X Y)) = ANDRIGHT (A AND B) (pi\<bullet>X) (pi\<bullet>Y)"
       
  2280 apply(auto simp add: perm_set_eq)
       
  2281 apply(rule_tac x="pi\<bullet>c" in exI)
       
  2282 apply(rule_tac x="pi\<bullet>a" in exI)
       
  2283 apply(rule_tac x="pi\<bullet>b" in exI)
       
  2284 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2285 apply(rule_tac x="pi\<bullet>N" in exI)
       
  2286 apply(simp)
       
  2287 apply(rule conjI)
       
  2288 apply(drule_tac pi="pi" in fic.eqvt(2))
       
  2289 apply(simp)
       
  2290 apply(rule conjI)
       
  2291 apply(rule_tac x="<a>:M" in exI)
       
  2292 apply(simp)
       
  2293 apply(rule_tac x="<b>:N" in exI)
       
  2294 apply(simp)
       
  2295 apply(rule_tac x="(rev pi)\<bullet>(<c>:AndR <a>.M <b>.N c)" in exI)
       
  2296 apply(perm_simp)
       
  2297 apply(rule_tac x="(rev pi)\<bullet>c" in exI)
       
  2298 apply(rule_tac x="(rev pi)\<bullet>a" in exI)
       
  2299 apply(rule_tac x="(rev pi)\<bullet>b" in exI)
       
  2300 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2301 apply(rule_tac x="(rev pi)\<bullet>N" in exI)
       
  2302 apply(simp)
       
  2303 apply(drule_tac pi="rev pi" in fic.eqvt(2))
       
  2304 apply(simp)
       
  2305 apply(drule sym)
       
  2306 apply(drule sym)
       
  2307 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2308 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2309 apply(simp)
       
  2310 done
       
  2311 
       
  2312 nominal_primrec
       
  2313   ANDLEFT1 :: "ty \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
       
  2314 where
       
  2315  "ANDLEFT1 (B AND C) X = { (y):AndL1 (x).M y | x y M. fin (AndL1 (x).M y) y \<and> (x):M \<in> X }"
       
  2316 apply(rule TrueI)+
       
  2317 done
       
  2318 
       
  2319 lemma ANDLEFT1_eqvt_name:
       
  2320   fixes pi::"name prm"
       
  2321   shows "(pi\<bullet>(ANDLEFT1 (A AND B) X)) = ANDLEFT1 (A AND B) (pi\<bullet>X)"
       
  2322 apply(auto simp add: perm_set_eq)
       
  2323 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2324 apply(rule_tac x="pi\<bullet>y" in exI) 
       
  2325 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2326 apply(simp)
       
  2327 apply(rule conjI)
       
  2328 apply(drule_tac pi="pi" in fin.eqvt(1))
       
  2329 apply(simp)
       
  2330 apply(rule_tac x="(xb):M" in exI)
       
  2331 apply(simp)
       
  2332 apply(rule_tac x="(rev pi)\<bullet>((y):AndL1 (xa).M y)" in exI)
       
  2333 apply(perm_simp)
       
  2334 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2335 apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
       
  2336 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2337 apply(simp)
       
  2338 apply(drule_tac pi="rev pi" in fin.eqvt(1))
       
  2339 apply(simp)
       
  2340 apply(drule sym)
       
  2341 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2342 apply(simp)
       
  2343 done
       
  2344 
       
  2345 lemma ANDLEFT1_eqvt_coname:
       
  2346   fixes pi::"coname prm"
       
  2347   shows "(pi\<bullet>(ANDLEFT1 (A AND B) X)) = ANDLEFT1 (A AND B) (pi\<bullet>X)"
       
  2348 apply(auto simp add: perm_set_eq)
       
  2349 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2350 apply(rule_tac x="pi\<bullet>y" in exI) 
       
  2351 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2352 apply(simp)
       
  2353 apply(rule conjI)
       
  2354 apply(drule_tac pi="pi" in fin.eqvt(2))
       
  2355 apply(simp)
       
  2356 apply(rule_tac x="(xb):M" in exI)
       
  2357 apply(simp)
       
  2358 apply(rule_tac x="(rev pi)\<bullet>((y):AndL1 (xa).M y)" in exI)
       
  2359 apply(perm_simp)
       
  2360 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2361 apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
       
  2362 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2363 apply(simp add: swap_simps)
       
  2364 apply(drule_tac pi="rev pi" in fin.eqvt(2))
       
  2365 apply(simp)
       
  2366 apply(drule sym)
       
  2367 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2368 apply(simp add: swap_simps)
       
  2369 done
       
  2370 
       
  2371 nominal_primrec
       
  2372   ANDLEFT2 :: "ty \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
       
  2373 where
       
  2374  "ANDLEFT2 (B AND C) X = { (y):AndL2 (x).M y | x y M. fin (AndL2 (x).M y) y \<and> (x):M \<in> X }"
       
  2375 apply(rule TrueI)+
       
  2376 done
       
  2377 
       
  2378 lemma ANDLEFT2_eqvt_name:
       
  2379   fixes pi::"name prm"
       
  2380   shows "(pi\<bullet>(ANDLEFT2 (A AND B) X)) = ANDLEFT2 (A AND B) (pi\<bullet>X)"
       
  2381 apply(auto simp add: perm_set_eq)
       
  2382 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2383 apply(rule_tac x="pi\<bullet>y" in exI) 
       
  2384 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2385 apply(simp)
       
  2386 apply(rule conjI)
       
  2387 apply(drule_tac pi="pi" in fin.eqvt(1))
       
  2388 apply(simp)
       
  2389 apply(rule_tac x="(xb):M" in exI)
       
  2390 apply(simp)
       
  2391 apply(rule_tac x="(rev pi)\<bullet>((y):AndL2 (xa).M y)" in exI)
       
  2392 apply(perm_simp)
       
  2393 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2394 apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
       
  2395 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2396 apply(simp)
       
  2397 apply(drule_tac pi="rev pi" in fin.eqvt(1))
       
  2398 apply(simp)
       
  2399 apply(drule sym)
       
  2400 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2401 apply(simp)
       
  2402 done
       
  2403 
       
  2404 lemma ANDLEFT2_eqvt_coname:
       
  2405   fixes pi::"coname prm"
       
  2406   shows "(pi\<bullet>(ANDLEFT2 (A AND B) X)) = ANDLEFT2 (A AND B) (pi\<bullet>X)"
       
  2407 apply(auto simp add: perm_set_eq)
       
  2408 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2409 apply(rule_tac x="pi\<bullet>y" in exI) 
       
  2410 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2411 apply(simp)
       
  2412 apply(rule conjI)
       
  2413 apply(drule_tac pi="pi" in fin.eqvt(2))
       
  2414 apply(simp)
       
  2415 apply(rule_tac x="(xb):M" in exI)
       
  2416 apply(simp)
       
  2417 apply(rule_tac x="(rev pi)\<bullet>((y):AndL2 (xa).M y)" in exI)
       
  2418 apply(perm_simp)
       
  2419 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2420 apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
       
  2421 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2422 apply(simp add: swap_simps)
       
  2423 apply(drule_tac pi="rev pi" in fin.eqvt(2))
       
  2424 apply(simp)
       
  2425 apply(drule sym)
       
  2426 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2427 apply(simp add: swap_simps)
       
  2428 done
       
  2429 
       
  2430 nominal_primrec
       
  2431   ORLEFT :: "ty \<Rightarrow> ntrm set \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
       
  2432 where
       
  2433  "ORLEFT (B OR C) X Y = 
       
  2434             { (z):OrL (x).M (y).N z | x y z M N. fin (OrL (x).M (y).N z) z \<and> (x):M \<in> X \<and> (y):N \<in> Y }"
       
  2435 apply(rule TrueI)+
       
  2436 done
       
  2437 
       
  2438 lemma ORLEFT_eqvt_name:
       
  2439   fixes pi::"name prm"
       
  2440   shows "(pi\<bullet>(ORLEFT (A OR B) X Y)) = ORLEFT (A OR B) (pi\<bullet>X) (pi\<bullet>Y)"
       
  2441 apply(auto simp add: perm_set_eq)
       
  2442 apply(rule_tac x="pi\<bullet>xb" in exI)
       
  2443 apply(rule_tac x="pi\<bullet>y" in exI)
       
  2444 apply(rule_tac x="pi\<bullet>z" in exI)
       
  2445 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2446 apply(rule_tac x="pi\<bullet>N" in exI)
       
  2447 apply(simp)
       
  2448 apply(rule conjI)
       
  2449 apply(drule_tac pi="pi" in fin.eqvt(1))
       
  2450 apply(simp)
       
  2451 apply(rule conjI)
       
  2452 apply(rule_tac x="(xb):M" in exI)
       
  2453 apply(simp)
       
  2454 apply(rule_tac x="(y):N" in exI)
       
  2455 apply(simp)
       
  2456 apply(rule_tac x="(rev pi)\<bullet>((z):OrL (xa).M (y).N z)" in exI)
       
  2457 apply(perm_simp)
       
  2458 apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
       
  2459 apply(rule_tac x="(rev pi)\<bullet>y" in exI)
       
  2460 apply(rule_tac x="(rev pi)\<bullet>z" in exI)
       
  2461 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2462 apply(rule_tac x="(rev pi)\<bullet>N" in exI)
       
  2463 apply(simp)
       
  2464 apply(drule_tac pi="rev pi" in fin.eqvt(1))
       
  2465 apply(simp)
       
  2466 apply(drule sym)
       
  2467 apply(drule sym)
       
  2468 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2469 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2470 apply(simp)
       
  2471 done
       
  2472 
       
  2473 lemma ORLEFT_eqvt_coname:
       
  2474   fixes pi::"coname prm"
       
  2475   shows "(pi\<bullet>(ORLEFT (A OR B) X Y)) = ORLEFT (A OR B) (pi\<bullet>X) (pi\<bullet>Y)"
       
  2476 apply(auto simp add: perm_set_eq)
       
  2477 apply(rule_tac x="pi\<bullet>xb" in exI)
       
  2478 apply(rule_tac x="pi\<bullet>y" in exI)
       
  2479 apply(rule_tac x="pi\<bullet>z" in exI)
       
  2480 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2481 apply(rule_tac x="pi\<bullet>N" in exI)
       
  2482 apply(simp)
       
  2483 apply(rule conjI)
       
  2484 apply(drule_tac pi="pi" in fin.eqvt(2))
       
  2485 apply(simp)
       
  2486 apply(rule conjI)
       
  2487 apply(rule_tac x="(xb):M" in exI)
       
  2488 apply(simp)
       
  2489 apply(rule_tac x="(y):N" in exI)
       
  2490 apply(simp)
       
  2491 apply(rule_tac x="(rev pi)\<bullet>((z):OrL (xa).M (y).N z)" in exI)
       
  2492 apply(perm_simp)
       
  2493 apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
       
  2494 apply(rule_tac x="(rev pi)\<bullet>y" in exI)
       
  2495 apply(rule_tac x="(rev pi)\<bullet>z" in exI)
       
  2496 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2497 apply(rule_tac x="(rev pi)\<bullet>N" in exI)
       
  2498 apply(simp add: swap_simps)
       
  2499 apply(drule_tac pi="rev pi" in fin.eqvt(2))
       
  2500 apply(simp)
       
  2501 apply(drule sym)
       
  2502 apply(drule sym)
       
  2503 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2504 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2505 apply(simp add: swap_simps)
       
  2506 done
       
  2507 
       
  2508 nominal_primrec
       
  2509   ORRIGHT1 :: "ty \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
       
  2510 where
       
  2511  "ORRIGHT1 (B OR C) X = { <b>:OrR1 <a>.M b | a b M. fic (OrR1 <a>.M b) b \<and> <a>:M \<in> X }"
       
  2512 apply(rule TrueI)+
       
  2513 done
       
  2514 
       
  2515 lemma ORRIGHT1_eqvt_name:
       
  2516   fixes pi::"name prm"
       
  2517   shows "(pi\<bullet>(ORRIGHT1 (A OR B) X)) = ORRIGHT1 (A OR B) (pi\<bullet>X)"
       
  2518 apply(auto simp add: perm_set_eq)
       
  2519 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2520 apply(rule_tac x="pi\<bullet>b" in exI) 
       
  2521 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2522 apply(simp)
       
  2523 apply(rule conjI)
       
  2524 apply(drule_tac pi="pi" in fic.eqvt(1))
       
  2525 apply(simp)
       
  2526 apply(rule_tac x="<a>:M" in exI)
       
  2527 apply(perm_simp)
       
  2528 apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR1 <a>.M b)" in exI)
       
  2529 apply(perm_simp)
       
  2530 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2531 apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
       
  2532 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2533 apply(simp add: swap_simps)
       
  2534 apply(drule_tac pi="rev pi" in fic.eqvt(1))
       
  2535 apply(simp)
       
  2536 apply(drule sym)
       
  2537 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2538 apply(simp add: swap_simps)
       
  2539 done
       
  2540 
       
  2541 lemma ORRIGHT1_eqvt_coname:
       
  2542   fixes pi::"coname prm"
       
  2543   shows "(pi\<bullet>(ORRIGHT1 (A OR B) X)) = ORRIGHT1 (A OR B) (pi\<bullet>X)"
       
  2544 apply(auto simp add: perm_set_eq)
       
  2545 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2546 apply(rule_tac x="pi\<bullet>b" in exI) 
       
  2547 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2548 apply(simp)
       
  2549 apply(rule conjI)
       
  2550 apply(drule_tac pi="pi" in fic.eqvt(2))
       
  2551 apply(simp)
       
  2552 apply(rule_tac x="<a>:M" in exI)
       
  2553 apply(perm_simp)
       
  2554 apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR1 <a>.M b)" in exI)
       
  2555 apply(perm_simp)
       
  2556 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2557 apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
       
  2558 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2559 apply(simp)
       
  2560 apply(drule_tac pi="rev pi" in fic.eqvt(2))
       
  2561 apply(simp)
       
  2562 apply(drule sym)
       
  2563 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2564 apply(simp)
       
  2565 done
       
  2566 
       
  2567 nominal_primrec
       
  2568   ORRIGHT2 :: "ty \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
       
  2569 where
       
  2570  "ORRIGHT2 (B OR C) X = { <b>:OrR2 <a>.M b | a b M. fic (OrR2 <a>.M b) b \<and> <a>:M \<in> X }"
       
  2571 apply(rule TrueI)+
       
  2572 done
       
  2573 
       
  2574 lemma ORRIGHT2_eqvt_name:
       
  2575   fixes pi::"name prm"
       
  2576   shows "(pi\<bullet>(ORRIGHT2 (A OR B) X)) = ORRIGHT2 (A OR B) (pi\<bullet>X)"
       
  2577 apply(auto simp add: perm_set_eq)
       
  2578 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2579 apply(rule_tac x="pi\<bullet>b" in exI) 
       
  2580 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2581 apply(simp)
       
  2582 apply(rule conjI)
       
  2583 apply(drule_tac pi="pi" in fic.eqvt(1))
       
  2584 apply(simp)
       
  2585 apply(rule_tac x="<a>:M" in exI)
       
  2586 apply(perm_simp)
       
  2587 apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR2 <a>.M b)" in exI)
       
  2588 apply(perm_simp)
       
  2589 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2590 apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
       
  2591 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2592 apply(simp add: swap_simps)
       
  2593 apply(drule_tac pi="rev pi" in fic.eqvt(1))
       
  2594 apply(simp)
       
  2595 apply(drule sym)
       
  2596 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2597 apply(simp add: swap_simps)
       
  2598 done
       
  2599 
       
  2600 lemma ORRIGHT2_eqvt_coname:
       
  2601   fixes pi::"coname prm"
       
  2602   shows "(pi\<bullet>(ORRIGHT2 (A OR B) X)) = ORRIGHT2 (A OR B) (pi\<bullet>X)"
       
  2603 apply(auto simp add: perm_set_eq)
       
  2604 apply(rule_tac x="pi\<bullet>a" in exI) 
       
  2605 apply(rule_tac x="pi\<bullet>b" in exI) 
       
  2606 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2607 apply(simp)
       
  2608 apply(rule conjI)
       
  2609 apply(drule_tac pi="pi" in fic.eqvt(2))
       
  2610 apply(simp)
       
  2611 apply(rule_tac x="<a>:M" in exI)
       
  2612 apply(perm_simp)
       
  2613 apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR2 <a>.M b)" in exI)
       
  2614 apply(perm_simp)
       
  2615 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2616 apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
       
  2617 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2618 apply(simp)
       
  2619 apply(drule_tac pi="rev pi" in fic.eqvt(2))
       
  2620 apply(simp)
       
  2621 apply(drule sym)
       
  2622 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2623 apply(simp)
       
  2624 done
       
  2625 
       
  2626 nominal_primrec
       
  2627   IMPRIGHT :: "ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set \<Rightarrow> ntrm set \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
       
  2628 where
       
  2629  "IMPRIGHT (B IMP C) X Y Z U= 
       
  2630         { <b>:ImpR (x).<a>.M b | x a b M. fic (ImpR (x).<a>.M b) b 
       
  2631                                         \<and> (\<forall>z P. x\<sharp>(z,P) \<and> (z):P \<in> Z \<longrightarrow> (x):(M{a:=(z).P}) \<in> X)
       
  2632                                         \<and> (\<forall>c Q. a\<sharp>(c,Q) \<and> <c>:Q \<in> U \<longrightarrow> <a>:(M{x:=<c>.Q}) \<in> Y)}"
       
  2633 apply(rule TrueI)+
       
  2634 done
       
  2635 
       
  2636 lemma IMPRIGHT_eqvt_name:
       
  2637   fixes pi::"name prm"
       
  2638   shows "(pi\<bullet>(IMPRIGHT (A IMP B) X Y Z U)) = IMPRIGHT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y) (pi\<bullet>Z) (pi\<bullet>U)"
       
  2639 apply(auto simp add: perm_set_eq)
       
  2640 apply(rule_tac x="pi\<bullet>xb" in exI)
       
  2641 apply(rule_tac x="pi\<bullet>a" in exI)
       
  2642 apply(rule_tac x="pi\<bullet>b" in exI)
       
  2643 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2644 apply(simp)
       
  2645 apply(rule conjI)
       
  2646 apply(drule_tac pi="pi" in fic.eqvt(1))
       
  2647 apply(simp)
       
  2648 apply(rule conjI)
       
  2649 apply(auto)[1]
       
  2650 apply(rule_tac x="(xb):(M{a:=((rev pi)\<bullet>z).((rev pi)\<bullet>P)})" in exI)
       
  2651 apply(perm_simp add: csubst_eqvt)
       
  2652 apply(drule sym)
       
  2653 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2654 apply(simp)
       
  2655 apply(simp add: fresh_right)
       
  2656 apply(auto)[1]
       
  2657 apply(rule_tac x="<a>:(M{xb:=<((rev pi)\<bullet>c)>.((rev pi)\<bullet>Q)})" in exI)
       
  2658 apply(perm_simp add: nsubst_eqvt)
       
  2659 apply(drule sym)
       
  2660 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2661 apply(simp add: swap_simps fresh_left)
       
  2662 apply(rule_tac x="(rev pi)\<bullet>(<b>:ImpR xa.<a>.M b)" in exI)
       
  2663 apply(perm_simp)
       
  2664 apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
       
  2665 apply(rule_tac x="(rev pi)\<bullet>a" in exI)
       
  2666 apply(rule_tac x="(rev pi)\<bullet>b" in exI)
       
  2667 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2668 apply(simp add: swap_simps)
       
  2669 apply(drule_tac pi="rev pi" in fic.eqvt(1))
       
  2670 apply(simp add: swap_simps)
       
  2671 apply(rule conjI)
       
  2672 apply(auto)[1]
       
  2673 apply(drule_tac x="pi\<bullet>z" in spec)
       
  2674 apply(drule_tac x="pi\<bullet>P" in spec)
       
  2675 apply(drule mp)
       
  2676 apply(simp add: fresh_right)
       
  2677 apply(rule_tac x="(z):P" in exI)
       
  2678 apply(simp)
       
  2679 apply(auto)[1]
       
  2680 apply(drule sym)
       
  2681 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2682 apply(perm_simp add: csubst_eqvt fresh_right)
       
  2683 apply(auto)[1]
       
  2684 apply(drule_tac x="pi\<bullet>c" in spec)
       
  2685 apply(drule_tac x="pi\<bullet>Q" in spec)
       
  2686 apply(drule mp)
       
  2687 apply(simp add: swap_simps fresh_left)
       
  2688 apply(rule_tac x="<c>:Q" in exI)
       
  2689 apply(simp add: swap_simps)
       
  2690 apply(auto)[1]
       
  2691 apply(drule sym)
       
  2692 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2693 apply(perm_simp add: nsubst_eqvt)
       
  2694 done
       
  2695 
       
  2696 lemma IMPRIGHT_eqvt_coname:
       
  2697   fixes pi::"coname prm"
       
  2698   shows "(pi\<bullet>(IMPRIGHT (A IMP B) X Y Z U)) = IMPRIGHT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y) (pi\<bullet>Z) (pi\<bullet>U)"
       
  2699 apply(auto simp add: perm_set_eq)
       
  2700 apply(rule_tac x="pi\<bullet>xb" in exI)
       
  2701 apply(rule_tac x="pi\<bullet>a" in exI)
       
  2702 apply(rule_tac x="pi\<bullet>b" in exI)
       
  2703 apply(rule_tac x="pi\<bullet>M" in exI)
       
  2704 apply(simp)
       
  2705 apply(rule conjI)
       
  2706 apply(drule_tac pi="pi" in fic.eqvt(2))
       
  2707 apply(simp)
       
  2708 apply(rule conjI)
       
  2709 apply(auto)[1]
       
  2710 apply(rule_tac x="(xb):(M{a:=((rev pi)\<bullet>z).((rev pi)\<bullet>P)})" in exI)
       
  2711 apply(perm_simp add: csubst_eqvt)
       
  2712 apply(drule sym)
       
  2713 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2714 apply(simp add: swap_simps fresh_left)
       
  2715 apply(auto)[1]
       
  2716 apply(rule_tac x="<a>:(M{xb:=<((rev pi)\<bullet>c)>.((rev pi)\<bullet>Q)})" in exI)
       
  2717 apply(perm_simp add: nsubst_eqvt)
       
  2718 apply(drule sym)
       
  2719 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2720 apply(simp add: fresh_right)
       
  2721 apply(rule_tac x="(rev pi)\<bullet>(<b>:ImpR xa.<a>.M b)" in exI)
       
  2722 apply(perm_simp)
       
  2723 apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
       
  2724 apply(rule_tac x="(rev pi)\<bullet>a" in exI)
       
  2725 apply(rule_tac x="(rev pi)\<bullet>b" in exI)
       
  2726 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2727 apply(simp add: swap_simps)
       
  2728 apply(drule_tac pi="rev pi" in fic.eqvt(2))
       
  2729 apply(simp add: swap_simps)
       
  2730 apply(rule conjI)
       
  2731 apply(auto)[1]
       
  2732 apply(drule_tac x="pi\<bullet>z" in spec)
       
  2733 apply(drule_tac x="pi\<bullet>P" in spec)
       
  2734 apply(simp add: swap_simps fresh_left)
       
  2735 apply(drule mp)
       
  2736 apply(rule_tac x="(z):P" in exI)
       
  2737 apply(simp add: swap_simps)
       
  2738 apply(auto)[1]
       
  2739 apply(drule sym)
       
  2740 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2741 apply(perm_simp add: csubst_eqvt fresh_right)
       
  2742 apply(auto)[1]
       
  2743 apply(drule_tac x="pi\<bullet>c" in spec)
       
  2744 apply(drule_tac x="pi\<bullet>Q" in spec)
       
  2745 apply(simp add: fresh_right)
       
  2746 apply(drule mp)
       
  2747 apply(rule_tac x="<c>:Q" in exI)
       
  2748 apply(simp)
       
  2749 apply(auto)[1]
       
  2750 apply(drule sym)
       
  2751 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2752 apply(perm_simp add: nsubst_eqvt fresh_right)
       
  2753 done
       
  2754 
       
  2755 nominal_primrec
       
  2756   IMPLEFT :: "ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
       
  2757 where
       
  2758  "IMPLEFT (B IMP C) X Y = 
       
  2759         { (y):ImpL <a>.M (x).N y | x a y M N. fin (ImpL <a>.M (x).N y) y \<and> <a>:M \<in> X \<and> (x):N \<in> Y }"
       
  2760 apply(rule TrueI)+
       
  2761 done
       
  2762 
       
  2763 lemma IMPLEFT_eqvt_name:
       
  2764   fixes pi::"name prm"
       
  2765   shows "(pi\<bullet>(IMPLEFT (A IMP B) X Y)) = IMPLEFT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y)"
       
  2766 apply(auto simp add: perm_set_eq)
       
  2767 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2768 apply(rule_tac x="pi\<bullet>a" in exI)
       
  2769 apply(rule_tac x="pi\<bullet>y" in exI) 
       
  2770 apply(rule_tac x="pi\<bullet>M" in exI) 
       
  2771 apply(rule_tac x="pi\<bullet>N" in exI)
       
  2772 apply(simp)
       
  2773 apply(rule conjI)
       
  2774 apply(drule_tac pi="pi" in fin.eqvt(1))
       
  2775 apply(simp)
       
  2776 apply(rule conjI)
       
  2777 apply(rule_tac x="<a>:M" in exI)
       
  2778 apply(simp)
       
  2779 apply(rule_tac x="(xb):N" in exI)
       
  2780 apply(simp)
       
  2781 apply(rule_tac x="(rev pi)\<bullet>((y):ImpL <a>.M (xa).N y)" in exI)
       
  2782 apply(perm_simp)
       
  2783 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2784 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2785 apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
       
  2786 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2787 apply(rule_tac x="(rev pi)\<bullet>N" in exI)
       
  2788 apply(simp add: swap_simps)
       
  2789 apply(drule_tac pi="rev pi" in fin.eqvt(1))
       
  2790 apply(simp)
       
  2791 apply(drule sym)
       
  2792 apply(drule sym)
       
  2793 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2794 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  2795 apply(simp add: swap_simps)
       
  2796 done
       
  2797 
       
  2798 lemma IMPLEFT_eqvt_coname:
       
  2799   fixes pi::"coname prm"
       
  2800   shows "(pi\<bullet>(IMPLEFT (A IMP B) X Y)) = IMPLEFT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y)"
       
  2801 apply(auto simp add: perm_set_eq)
       
  2802 apply(rule_tac x="pi\<bullet>xb" in exI) 
       
  2803 apply(rule_tac x="pi\<bullet>a" in exI)
       
  2804 apply(rule_tac x="pi\<bullet>y" in exI) 
       
  2805 apply(rule_tac x="pi\<bullet>M" in exI) 
       
  2806 apply(rule_tac x="pi\<bullet>N" in exI)
       
  2807 apply(simp)
       
  2808 apply(rule conjI)
       
  2809 apply(drule_tac pi="pi" in fin.eqvt(2))
       
  2810 apply(simp)
       
  2811 apply(rule conjI)
       
  2812 apply(rule_tac x="<a>:M" in exI)
       
  2813 apply(simp)
       
  2814 apply(rule_tac x="(xb):N" in exI)
       
  2815 apply(simp)
       
  2816 apply(rule_tac x="(rev pi)\<bullet>((y):ImpL <a>.M (xa).N y)" in exI)
       
  2817 apply(perm_simp)
       
  2818 apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
       
  2819 apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
       
  2820 apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
       
  2821 apply(rule_tac x="(rev pi)\<bullet>M" in exI)
       
  2822 apply(rule_tac x="(rev pi)\<bullet>N" in exI)
       
  2823 apply(simp add: swap_simps)
       
  2824 apply(drule_tac pi="rev pi" in fin.eqvt(2))
       
  2825 apply(simp)
       
  2826 apply(drule sym)
       
  2827 apply(drule sym)
       
  2828 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2829 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2830 apply(simp add: swap_simps)
       
  2831 done
       
  2832 
       
  2833 lemma sum_cases:
       
  2834  shows "(\<exists>y. x=Inl y) \<or> (\<exists>y. x=Inr y)"
       
  2835 apply(rule_tac s="x" in sumE)
       
  2836 apply(auto)
       
  2837 done
       
  2838 
       
  2839 function
       
  2840   NEGc::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set"
       
  2841 and
       
  2842   NEGn::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set"
       
  2843 where
       
  2844   "NEGc (PR A)    X = AXIOMSc (PR A) \<union> BINDINGc (PR A) X"  
       
  2845 | "NEGc (NOT C)   X = AXIOMSc (NOT C) \<union> BINDINGc (NOT C) X 
       
  2846                          \<union> NOTRIGHT (NOT C) (lfp (NEGn C \<circ> NEGc C))"  
       
  2847 | "NEGc (C AND D) X = AXIOMSc (C AND D) \<union> BINDINGc (C AND D) X 
       
  2848                      \<union> ANDRIGHT (C AND D) (NEGc C (lfp (NEGn C \<circ> NEGc C))) (NEGc D (lfp (NEGn D \<circ> NEGc D)))"
       
  2849 | "NEGc (C OR D)  X = AXIOMSc (C OR D) \<union> BINDINGc (C OR D) X  
       
  2850                          \<union> ORRIGHT1 (C OR D) (NEGc C (lfp (NEGn C \<circ> NEGc C))) 
       
  2851                          \<union> ORRIGHT2 (C OR D) (NEGc D (lfp (NEGn D \<circ> NEGc D)))"
       
  2852 | "NEGc (C IMP D) X = AXIOMSc (C IMP D) \<union> BINDINGc (C IMP D) X 
       
  2853     \<union> IMPRIGHT (C IMP D) (lfp (NEGn C \<circ> NEGc C)) (NEGc D (lfp (NEGn D \<circ> NEGc D))) 
       
  2854                           (lfp (NEGn D \<circ> NEGc D)) (NEGc C (lfp (NEGn C \<circ> NEGc C)))"
       
  2855 | "NEGn (PR A)    X = AXIOMSn (PR A) \<union> BINDINGn (PR A) X"   
       
  2856 | "NEGn (NOT C)   X = AXIOMSn (NOT C) \<union> BINDINGn (NOT C) X 
       
  2857                          \<union> NOTLEFT (NOT C) (NEGc C (lfp (NEGn C \<circ> NEGc C)))"  
       
  2858 | "NEGn (C AND D) X = AXIOMSn (C AND D) \<union> BINDINGn (C AND D) X 
       
  2859                          \<union> ANDLEFT1 (C AND D) (lfp (NEGn C \<circ> NEGc C)) 
       
  2860                          \<union> ANDLEFT2 (C AND D) (lfp (NEGn D \<circ> NEGc D))"
       
  2861 | "NEGn (C OR D)  X = AXIOMSn (C OR D) \<union> BINDINGn (C OR D) X 
       
  2862                          \<union> ORLEFT (C OR D) (lfp (NEGn C \<circ> NEGc C)) (lfp (NEGn D \<circ> NEGc D))"
       
  2863 | "NEGn (C IMP D) X = AXIOMSn (C IMP D) \<union> BINDINGn (C IMP D) X 
       
  2864                          \<union> IMPLEFT (C IMP D) (NEGc C (lfp (NEGn C \<circ> NEGc C))) (lfp (NEGn D \<circ> NEGc D))"
       
  2865 using ty_cases sum_cases 
       
  2866 apply(auto simp add: ty.inject)
       
  2867 apply(drule_tac x="x" in meta_spec)
       
  2868 apply(auto simp add: ty.inject)
       
  2869 apply(rotate_tac 10)
       
  2870 apply(drule_tac x="a" in meta_spec)
       
  2871 apply(auto simp add: ty.inject)
       
  2872 apply(blast)
       
  2873 apply(blast)
       
  2874 apply(blast)
       
  2875 apply(rotate_tac 10)
       
  2876 apply(drule_tac x="a" in meta_spec)
       
  2877 apply(auto simp add: ty.inject)
       
  2878 apply(blast)
       
  2879 apply(blast)
       
  2880 apply(blast)
       
  2881 done
       
  2882 
       
  2883 termination
       
  2884 apply(relation "measure (sum_case (size\<circ>fst) (size\<circ>fst))")
       
  2885 apply(simp_all)
       
  2886 done
       
  2887 
       
  2888 text {* Candidates *}
       
  2889 
       
  2890 lemma test1:
       
  2891   shows "x\<in>(X\<union>Y) = (x\<in>X \<or> x\<in>Y)"
       
  2892 by blast
       
  2893 
       
  2894 lemma test2:
       
  2895   shows "x\<in>(X\<inter>Y) = (x\<in>X \<and> x\<in>Y)"
       
  2896 by blast
       
  2897 
       
  2898 lemma big_inter_eqvt:
       
  2899   fixes pi1::"name prm"
       
  2900   and   X::"('a::pt_name) set set"
       
  2901   and   pi2::"coname prm"
       
  2902   and   Y::"('b::pt_coname) set set"
       
  2903   shows "(pi1\<bullet>(\<Inter> X)) = \<Inter> (pi1\<bullet>X)"
       
  2904   and   "(pi2\<bullet>(\<Inter> Y)) = \<Inter> (pi2\<bullet>Y)"
       
  2905 apply(auto simp add: perm_set_eq)
       
  2906 apply(rule_tac x="(rev pi1)\<bullet>x" in exI)
       
  2907 apply(perm_simp)
       
  2908 apply(rule ballI)
       
  2909 apply(drule_tac x="pi1\<bullet>xa" in spec)
       
  2910 apply(auto)
       
  2911 apply(drule_tac x="xa" in spec)
       
  2912 apply(auto)[1]
       
  2913 apply(rule_tac x="(rev pi1)\<bullet>xb" in exI)
       
  2914 apply(perm_simp)
       
  2915 apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
       
  2916 apply(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst])
       
  2917 apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
       
  2918 apply(rule_tac x="(rev pi2)\<bullet>x" in exI)
       
  2919 apply(perm_simp)
       
  2920 apply(rule ballI)
       
  2921 apply(drule_tac x="pi2\<bullet>xa" in spec)
       
  2922 apply(auto)
       
  2923 apply(drule_tac x="xa" in spec)
       
  2924 apply(auto)[1]
       
  2925 apply(rule_tac x="(rev pi2)\<bullet>xb" in exI)
       
  2926 apply(perm_simp)
       
  2927 apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2928 apply(simp add: pt_set_bij[OF pt_coname_inst, OF at_coname_inst])
       
  2929 apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  2930 done
       
  2931 
       
  2932 lemma lfp_eqvt:
       
  2933   fixes pi1::"name prm"
       
  2934   and   f::"'a set \<Rightarrow> ('a::pt_name) set"
       
  2935   and   pi2::"coname prm"
       
  2936   and   g::"'b set \<Rightarrow> ('b::pt_coname) set"
       
  2937   shows "pi1\<bullet>(lfp f) = lfp (pi1\<bullet>f)"
       
  2938   and   "pi2\<bullet>(lfp g) = lfp (pi2\<bullet>g)"
       
  2939 apply(simp add: lfp_def)
       
  2940 apply(simp add: big_inter_eqvt)
       
  2941 apply(simp add: pt_Collect_eqvt[OF pt_name_inst, OF at_name_inst])
       
  2942 apply(subgoal_tac "{u. (pi1\<bullet>f) u \<subseteq> u} = {u. ((rev pi1)\<bullet>((pi1\<bullet>f) u)) \<subseteq> ((rev pi1)\<bullet>u)}")
       
  2943 apply(perm_simp)
       
  2944 apply(rule Collect_cong)
       
  2945 apply(rule iffI)
       
  2946 apply(rule subseteq_eqvt(1)[THEN iffD1])
       
  2947 apply(simp add: perm_bool)
       
  2948 apply(drule subseteq_eqvt(1)[THEN iffD2])
       
  2949 apply(simp add: perm_bool)
       
  2950 apply(simp add: lfp_def)
       
  2951 apply(simp add: big_inter_eqvt)
       
  2952 apply(simp add: pt_Collect_eqvt[OF pt_coname_inst, OF at_coname_inst])
       
  2953 apply(subgoal_tac "{u. (pi2\<bullet>g) u \<subseteq> u} = {u. ((rev pi2)\<bullet>((pi2\<bullet>g) u)) \<subseteq> ((rev pi2)\<bullet>u)}")
       
  2954 apply(perm_simp)
       
  2955 apply(rule Collect_cong)
       
  2956 apply(rule iffI)
       
  2957 apply(rule subseteq_eqvt(2)[THEN iffD1])
       
  2958 apply(simp add: perm_bool)
       
  2959 apply(drule subseteq_eqvt(2)[THEN iffD2])
       
  2960 apply(simp add: perm_bool)
       
  2961 done
       
  2962 
       
  2963 abbreviation
       
  2964   CANDn::"ty \<Rightarrow> ntrm set"  ("\<parallel>'(_')\<parallel>" [60] 60) 
       
  2965 where
       
  2966   "\<parallel>(B)\<parallel> \<equiv> lfp (NEGn B \<circ> NEGc B)" 
       
  2967 
       
  2968 abbreviation
       
  2969   CANDc::"ty \<Rightarrow> ctrm set"  ("\<parallel><_>\<parallel>" [60] 60)
       
  2970 where
       
  2971   "\<parallel><B>\<parallel> \<equiv> NEGc B (\<parallel>(B)\<parallel>)"
       
  2972 
       
  2973 lemma NEGn_decreasing:
       
  2974   shows "X\<subseteq>Y \<Longrightarrow> NEGn B Y \<subseteq> NEGn B X"
       
  2975 by (nominal_induct B rule: ty.strong_induct)
       
  2976    (auto dest: BINDINGn_decreasing)
       
  2977 
       
  2978 lemma NEGc_decreasing:
       
  2979   shows "X\<subseteq>Y \<Longrightarrow> NEGc B Y \<subseteq> NEGc B X"
       
  2980 by (nominal_induct B rule: ty.strong_induct)
       
  2981    (auto dest: BINDINGc_decreasing)
       
  2982 
       
  2983 lemma mono_NEGn_NEGc:
       
  2984   shows "mono (NEGn B \<circ> NEGc B)"
       
  2985   and   "mono (NEGc B \<circ> NEGn B)"
       
  2986 proof -
       
  2987   have "\<forall>X Y. X\<subseteq>Y \<longrightarrow> NEGn B (NEGc B X) \<subseteq> NEGn B (NEGc B Y)"
       
  2988   proof (intro strip)
       
  2989     fix X::"ntrm set" and Y::"ntrm set"
       
  2990     assume "X\<subseteq>Y"
       
  2991     then have "NEGc B Y \<subseteq> NEGc B X" by (simp add: NEGc_decreasing)
       
  2992     then show "NEGn B (NEGc B X) \<subseteq> NEGn B (NEGc B Y)" by (simp add: NEGn_decreasing)
       
  2993   qed
       
  2994   then show "mono (NEGn B \<circ> NEGc B)" by (simp add: mono_def)
       
  2995 next
       
  2996   have "\<forall>X Y. X\<subseteq>Y \<longrightarrow> NEGc B (NEGn B X) \<subseteq> NEGc B (NEGn B Y)"
       
  2997   proof (intro strip)
       
  2998     fix X::"ctrm set" and Y::"ctrm set"
       
  2999     assume "X\<subseteq>Y"
       
  3000     then have "NEGn B Y \<subseteq> NEGn B X" by (simp add: NEGn_decreasing)
       
  3001     then show "NEGc B (NEGn B X) \<subseteq> NEGc B (NEGn B Y)" by (simp add: NEGc_decreasing)
       
  3002   qed
       
  3003   then show "mono (NEGc B \<circ> NEGn B)" by (simp add: mono_def)
       
  3004 qed
       
  3005 
       
  3006 lemma NEG_simp:
       
  3007   shows "\<parallel><B>\<parallel> = NEGc B (\<parallel>(B)\<parallel>)"
       
  3008   and   "\<parallel>(B)\<parallel> = NEGn B (\<parallel><B>\<parallel>)"
       
  3009 proof -
       
  3010   show "\<parallel><B>\<parallel> = NEGc B (\<parallel>(B)\<parallel>)" by simp
       
  3011 next
       
  3012   have "\<parallel>(B)\<parallel> \<equiv> lfp (NEGn B \<circ> NEGc B)" by simp
       
  3013   then have "\<parallel>(B)\<parallel> = (NEGn B \<circ> NEGc B) (\<parallel>(B)\<parallel>)" using mono_NEGn_NEGc def_lfp_unfold by blast
       
  3014   then show "\<parallel>(B)\<parallel> = NEGn B (\<parallel><B>\<parallel>)" by simp
       
  3015 qed
       
  3016 
       
  3017 lemma NEG_elim:
       
  3018   shows "M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<in> NEGc B (\<parallel>(B)\<parallel>)"
       
  3019   and   "N \<in> \<parallel>(B)\<parallel> \<Longrightarrow> N \<in> NEGn B (\<parallel><B>\<parallel>)"
       
  3020 using NEG_simp by (blast)+
       
  3021 
       
  3022 lemma NEG_intro:
       
  3023   shows "M \<in> NEGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> M \<in> \<parallel><B>\<parallel>"
       
  3024   and   "N \<in> NEGn B (\<parallel><B>\<parallel>) \<Longrightarrow> N \<in> \<parallel>(B)\<parallel>"
       
  3025 using NEG_simp by (blast)+
       
  3026 
       
  3027 lemma NEGc_simps:
       
  3028   shows "NEGc (PR A) (\<parallel>(PR A)\<parallel>) = AXIOMSc (PR A) \<union> BINDINGc (PR A) (\<parallel>(PR A)\<parallel>)"  
       
  3029   and   "NEGc (NOT C) (\<parallel>(NOT C)\<parallel>) = AXIOMSc (NOT C) \<union> BINDINGc (NOT C) (\<parallel>(NOT C)\<parallel>) 
       
  3030                                         \<union> (NOTRIGHT (NOT C) (\<parallel>(C)\<parallel>))"  
       
  3031   and   "NEGc (C AND D) (\<parallel>(C AND D)\<parallel>) = AXIOMSc (C AND D) \<union> BINDINGc (C AND D) (\<parallel>(C AND D)\<parallel>) 
       
  3032                                         \<union> (ANDRIGHT (C AND D) (\<parallel><C>\<parallel>) (\<parallel><D>\<parallel>))"
       
  3033   and   "NEGc (C OR D) (\<parallel>(C OR D)\<parallel>) = AXIOMSc (C OR D) \<union> BINDINGc (C OR D)  (\<parallel>(C OR D)\<parallel>)
       
  3034                                         \<union> (ORRIGHT1 (C OR D) (\<parallel><C>\<parallel>)) \<union> (ORRIGHT2 (C OR D) (\<parallel><D>\<parallel>))"
       
  3035   and   "NEGc (C IMP D) (\<parallel>(C IMP D)\<parallel>) = AXIOMSc (C IMP D) \<union> BINDINGc (C IMP D) (\<parallel>(C IMP D)\<parallel>) 
       
  3036           \<union> (IMPRIGHT (C IMP D) (\<parallel>(C)\<parallel>) (\<parallel><D>\<parallel>) (\<parallel>(D)\<parallel>) (\<parallel><C>\<parallel>))"
       
  3037 by (simp_all only: NEGc.simps)
       
  3038 
       
  3039 lemma AXIOMS_in_CANDs:
       
  3040   shows "AXIOMSn B \<subseteq> (\<parallel>(B)\<parallel>)"
       
  3041   and   "AXIOMSc B \<subseteq> (\<parallel><B>\<parallel>)"
       
  3042 proof -
       
  3043   have "AXIOMSn B \<subseteq> NEGn B (\<parallel><B>\<parallel>)"
       
  3044     by (nominal_induct B rule: ty.strong_induct) (auto)
       
  3045   then show "AXIOMSn B \<subseteq> \<parallel>(B)\<parallel>" using NEG_simp by blast
       
  3046 next
       
  3047   have "AXIOMSc B \<subseteq> NEGc B (\<parallel>(B)\<parallel>)"
       
  3048     by (nominal_induct B rule: ty.strong_induct) (auto)
       
  3049   then show "AXIOMSc B \<subseteq> \<parallel><B>\<parallel>" using NEG_simp by blast
       
  3050 qed
       
  3051 
       
  3052 lemma Ax_in_CANDs:
       
  3053   shows "(y):Ax x a \<in> \<parallel>(B)\<parallel>"
       
  3054   and   "<b>:Ax x a \<in> \<parallel><B>\<parallel>"
       
  3055 proof -
       
  3056   have "(y):Ax x a \<in> AXIOMSn B" by (auto simp add: AXIOMSn_def)
       
  3057   also have "AXIOMSn B \<subseteq> \<parallel>(B)\<parallel>" by (rule AXIOMS_in_CANDs)
       
  3058   finally show "(y):Ax x a \<in> \<parallel>(B)\<parallel>" by simp
       
  3059 next
       
  3060   have "<b>:Ax x a \<in> AXIOMSc B" by (auto simp add: AXIOMSc_def)
       
  3061   also have "AXIOMSc B \<subseteq> \<parallel><B>\<parallel>" by (rule AXIOMS_in_CANDs)
       
  3062   finally show "<b>:Ax x a \<in> \<parallel><B>\<parallel>" by simp
       
  3063 qed
       
  3064 
       
  3065 lemma AXIOMS_eqvt_aux_name:
       
  3066   fixes pi::"name prm"
       
  3067   shows "M \<in> AXIOMSn B \<Longrightarrow> (pi\<bullet>M) \<in> AXIOMSn B" 
       
  3068   and   "N \<in> AXIOMSc B \<Longrightarrow> (pi\<bullet>N) \<in> AXIOMSc B"
       
  3069 apply(auto simp add: AXIOMSn_def AXIOMSc_def)
       
  3070 apply(rule_tac x="pi\<bullet>x" in exI)
       
  3071 apply(rule_tac x="pi\<bullet>y" in exI)
       
  3072 apply(rule_tac x="pi\<bullet>b" in exI)
       
  3073 apply(simp)
       
  3074 apply(rule_tac x="pi\<bullet>a" in exI)
       
  3075 apply(rule_tac x="pi\<bullet>y" in exI)
       
  3076 apply(rule_tac x="pi\<bullet>b" in exI)
       
  3077 apply(simp)
       
  3078 done
       
  3079 
       
  3080 lemma AXIOMS_eqvt_aux_coname:
       
  3081   fixes pi::"coname prm"
       
  3082   shows "M \<in> AXIOMSn B \<Longrightarrow> (pi\<bullet>M) \<in> AXIOMSn B" 
       
  3083   and   "N \<in> AXIOMSc B \<Longrightarrow> (pi\<bullet>N) \<in> AXIOMSc B"
       
  3084 apply(auto simp add: AXIOMSn_def AXIOMSc_def)
       
  3085 apply(rule_tac x="pi\<bullet>x" in exI)
       
  3086 apply(rule_tac x="pi\<bullet>y" in exI)
       
  3087 apply(rule_tac x="pi\<bullet>b" in exI)
       
  3088 apply(simp)
       
  3089 apply(rule_tac x="pi\<bullet>a" in exI)
       
  3090 apply(rule_tac x="pi\<bullet>y" in exI)
       
  3091 apply(rule_tac x="pi\<bullet>b" in exI)
       
  3092 apply(simp)
       
  3093 done
       
  3094 
       
  3095 lemma AXIOMS_eqvt_name:
       
  3096   fixes pi::"name prm"
       
  3097   shows "(pi\<bullet>AXIOMSn B) = AXIOMSn B" 
       
  3098   and   "(pi\<bullet>AXIOMSc B) = AXIOMSc B"
       
  3099 apply(auto)
       
  3100 apply(simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
       
  3101 apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_name(1))
       
  3102 apply(perm_simp)
       
  3103 apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_name(1))
       
  3104 apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
       
  3105 apply(simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
       
  3106 apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_name(2))
       
  3107 apply(perm_simp)
       
  3108 apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_name(2))
       
  3109 apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
       
  3110 done
       
  3111 
       
  3112 lemma AXIOMS_eqvt_coname:
       
  3113   fixes pi::"coname prm"
       
  3114   shows "(pi\<bullet>AXIOMSn B) = AXIOMSn B" 
       
  3115   and   "(pi\<bullet>AXIOMSc B) = AXIOMSc B"
       
  3116 apply(auto)
       
  3117 apply(simp add: pt_set_bij1a[OF pt_coname_inst, OF at_coname_inst])
       
  3118 apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_coname(1))
       
  3119 apply(perm_simp)
       
  3120 apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_coname(1))
       
  3121 apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  3122 apply(simp add: pt_set_bij1a[OF pt_coname_inst, OF at_coname_inst])
       
  3123 apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_coname(2))
       
  3124 apply(perm_simp)
       
  3125 apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_coname(2))
       
  3126 apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  3127 done
       
  3128 
       
  3129 lemma BINDING_eqvt_name:
       
  3130   fixes pi::"name prm"
       
  3131   shows "(pi\<bullet>(BINDINGn B X)) = BINDINGn B (pi\<bullet>X)" 
       
  3132   and   "(pi\<bullet>(BINDINGc B Y)) = BINDINGc B (pi\<bullet>Y)" 
       
  3133 apply(auto simp add: BINDINGn_def BINDINGc_def perm_set_eq)
       
  3134 apply(rule_tac x="pi\<bullet>xb" in exI)
       
  3135 apply(rule_tac x="pi\<bullet>M" in exI)
       
  3136 apply(simp)
       
  3137 apply(auto)[1]
       
  3138 apply(drule_tac x="(rev pi)\<bullet>a" in spec)
       
  3139 apply(drule_tac x="(rev pi)\<bullet>P" in spec)
       
  3140 apply(drule mp)
       
  3141 apply(drule sym)
       
  3142 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  3143 apply(simp)
       
  3144 apply(drule_tac ?pi1.0="pi" in SNa_eqvt(1))
       
  3145 apply(perm_simp add: nsubst_eqvt)
       
  3146 apply(rule_tac x="(rev pi\<bullet>xa):(rev pi\<bullet>M)" in exI)
       
  3147 apply(perm_simp)
       
  3148 apply(rule_tac x="rev pi\<bullet>xa" in exI)
       
  3149 apply(rule_tac x="rev pi\<bullet>M" in exI)
       
  3150 apply(simp)
       
  3151 apply(auto)[1]
       
  3152 apply(drule_tac x="pi\<bullet>a" in spec)
       
  3153 apply(drule_tac x="pi\<bullet>P" in spec)
       
  3154 apply(drule mp)
       
  3155 apply(force)
       
  3156 apply(drule_tac ?pi1.0="rev pi" in SNa_eqvt(1))
       
  3157 apply(perm_simp add: nsubst_eqvt)
       
  3158 apply(rule_tac x="pi\<bullet>a" in exI)
       
  3159 apply(rule_tac x="pi\<bullet>M" in exI)
       
  3160 apply(simp)
       
  3161 apply(auto)[1]
       
  3162 apply(drule_tac x="(rev pi)\<bullet>x" in spec)
       
  3163 apply(drule_tac x="(rev pi)\<bullet>P" in spec)
       
  3164 apply(drule mp)
       
  3165 apply(drule sym)
       
  3166 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
       
  3167 apply(simp)
       
  3168 apply(drule_tac ?pi1.0="pi" in SNa_eqvt(1))
       
  3169 apply(perm_simp add: csubst_eqvt)
       
  3170 apply(rule_tac x="<(rev pi\<bullet>a)>:(rev pi\<bullet>M)" in exI)
       
  3171 apply(perm_simp)
       
  3172 apply(rule_tac x="rev pi\<bullet>a" in exI)
       
  3173 apply(rule_tac x="rev pi\<bullet>M" in exI)
       
  3174 apply(simp add: swap_simps)
       
  3175 apply(auto)[1]
       
  3176 apply(drule_tac x="pi\<bullet>x" in spec)
       
  3177 apply(drule_tac x="pi\<bullet>P" in spec)
       
  3178 apply(drule mp)
       
  3179 apply(force)
       
  3180 apply(drule_tac ?pi1.0="rev pi" in SNa_eqvt(1))
       
  3181 apply(perm_simp add: csubst_eqvt)
       
  3182 done
       
  3183 
       
  3184 lemma BINDING_eqvt_coname:
       
  3185   fixes pi::"coname prm"
       
  3186   shows "(pi\<bullet>(BINDINGn B X)) = BINDINGn B (pi\<bullet>X)" 
       
  3187   and   "(pi\<bullet>(BINDINGc B Y)) = BINDINGc B (pi\<bullet>Y)" 
       
  3188 apply(auto simp add: BINDINGn_def BINDINGc_def perm_set_eq)
       
  3189 apply(rule_tac x="pi\<bullet>xb" in exI)
       
  3190 apply(rule_tac x="pi\<bullet>M" in exI)
       
  3191 apply(simp)
       
  3192 apply(auto)[1]
       
  3193 apply(drule_tac x="(rev pi)\<bullet>a" in spec)
       
  3194 apply(drule_tac x="(rev pi)\<bullet>P" in spec)
       
  3195 apply(drule mp)
       
  3196 apply(drule sym)
       
  3197 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  3198 apply(simp)
       
  3199 apply(drule_tac ?pi2.0="pi" in SNa_eqvt(2))
       
  3200 apply(perm_simp add: nsubst_eqvt)
       
  3201 apply(rule_tac x="(rev pi\<bullet>xa):(rev pi\<bullet>M)" in exI)
       
  3202 apply(perm_simp)
       
  3203 apply(rule_tac x="rev pi\<bullet>xa" in exI)
       
  3204 apply(rule_tac x="rev pi\<bullet>M" in exI)
       
  3205 apply(simp add: swap_simps)
       
  3206 apply(auto)[1]
       
  3207 apply(drule_tac x="pi\<bullet>a" in spec)
       
  3208 apply(drule_tac x="pi\<bullet>P" in spec)
       
  3209 apply(drule mp)
       
  3210 apply(force)
       
  3211 apply(drule_tac ?pi2.0="rev pi" in SNa_eqvt(2))
       
  3212 apply(perm_simp add: nsubst_eqvt)
       
  3213 apply(rule_tac x="pi\<bullet>a" in exI)
       
  3214 apply(rule_tac x="pi\<bullet>M" in exI)
       
  3215 apply(simp)
       
  3216 apply(auto)[1]
       
  3217 apply(drule_tac x="(rev pi)\<bullet>x" in spec)
       
  3218 apply(drule_tac x="(rev pi)\<bullet>P" in spec)
       
  3219 apply(drule mp)
       
  3220 apply(drule sym)
       
  3221 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
       
  3222 apply(simp)
       
  3223 apply(drule_tac ?pi2.0="pi" in SNa_eqvt(2))
       
  3224 apply(perm_simp add: csubst_eqvt)
       
  3225 apply(rule_tac x="<(rev pi\<bullet>a)>:(rev pi\<bullet>M)" in exI)
       
  3226 apply(perm_simp)
       
  3227 apply(rule_tac x="rev pi\<bullet>a" in exI)
       
  3228 apply(rule_tac x="rev pi\<bullet>M" in exI)
       
  3229 apply(simp)
       
  3230 apply(auto)[1]
       
  3231 apply(drule_tac x="pi\<bullet>x" in spec)
       
  3232 apply(drule_tac x="pi\<bullet>P" in spec)
       
  3233 apply(drule mp)
       
  3234 apply(force)
       
  3235 apply(drule_tac ?pi2.0="rev pi" in SNa_eqvt(2))
       
  3236 apply(perm_simp add: csubst_eqvt)
       
  3237 done
       
  3238 
       
  3239 lemma CAND_eqvt_name:
       
  3240   fixes pi::"name prm"
       
  3241   shows   "(pi\<bullet>(\<parallel>(B)\<parallel>)) = (\<parallel>(B)\<parallel>)"
       
  3242   and     "(pi\<bullet>(\<parallel><B>\<parallel>)) = (\<parallel><B>\<parallel>)"
       
  3243 proof (nominal_induct B rule: ty.strong_induct)
       
  3244   case (PR X)
       
  3245   { case 1 show ?case 
       
  3246       apply -
       
  3247       apply(simp add: lfp_eqvt)
       
  3248       apply(simp add: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3249       apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name)
       
  3250       apply(perm_simp)
       
  3251     done
       
  3252   next
       
  3253     case 2 show ?case
       
  3254       apply -
       
  3255       apply(simp only: NEGc_simps)
       
  3256       apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name)
       
  3257       apply(simp add: lfp_eqvt)
       
  3258       apply(simp add: comp_def)
       
  3259       apply(simp add: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3260       apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name)
       
  3261       apply(perm_simp)
       
  3262       done
       
  3263   }
       
  3264 next
       
  3265   case (NOT B)
       
  3266   have ih1: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3267   have ih2: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3268   have g: "pi\<bullet>(\<parallel>(NOT B)\<parallel>) = (\<parallel>(NOT B)\<parallel>)"
       
  3269     apply -
       
  3270     apply(simp only: lfp_eqvt)
       
  3271     apply(simp only: comp_def)
       
  3272     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3273     apply(simp only: NEGc.simps NEGn.simps)
       
  3274     apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name NOTRIGHT_eqvt_name NOTLEFT_eqvt_name)
       
  3275     apply(perm_simp add: ih1 ih2)
       
  3276     done
       
  3277   { case 1 show ?case by (rule g)
       
  3278   next 
       
  3279     case 2 show ?case
       
  3280       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name NOTRIGHT_eqvt_name ih1 ih2 g)
       
  3281   }
       
  3282 next
       
  3283   case (AND A B)
       
  3284   have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
       
  3285   have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
       
  3286   have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3287   have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3288   have g: "pi\<bullet>(\<parallel>(A AND B)\<parallel>) = (\<parallel>(A AND B)\<parallel>)"
       
  3289     apply -
       
  3290     apply(simp only: lfp_eqvt)
       
  3291     apply(simp only: comp_def)
       
  3292     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3293     apply(simp only: NEGc.simps NEGn.simps)
       
  3294     apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name ANDRIGHT_eqvt_name 
       
  3295                      ANDLEFT2_eqvt_name ANDLEFT1_eqvt_name)
       
  3296     apply(perm_simp add: ih1 ih2 ih3 ih4)
       
  3297     done
       
  3298   { case 1 show ?case by (rule g)
       
  3299   next 
       
  3300     case 2 show ?case
       
  3301       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name 
       
  3302                      ANDRIGHT_eqvt_name ANDLEFT1_eqvt_name ANDLEFT2_eqvt_name ih1 ih2 ih3 ih4 g)
       
  3303   }
       
  3304 next
       
  3305   case (OR A B)
       
  3306   have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
       
  3307   have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
       
  3308   have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3309   have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3310   have g: "pi\<bullet>(\<parallel>(A OR B)\<parallel>) = (\<parallel>(A OR B)\<parallel>)"
       
  3311     apply -
       
  3312     apply(simp only: lfp_eqvt)
       
  3313     apply(simp only: comp_def)
       
  3314     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3315     apply(simp only: NEGc.simps NEGn.simps)
       
  3316     apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name ORRIGHT1_eqvt_name 
       
  3317                      ORRIGHT2_eqvt_name ORLEFT_eqvt_name)
       
  3318     apply(perm_simp add: ih1 ih2 ih3 ih4)
       
  3319     done
       
  3320   { case 1 show ?case by (rule g)
       
  3321   next 
       
  3322     case 2 show ?case
       
  3323       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name 
       
  3324                      ORRIGHT1_eqvt_name ORRIGHT2_eqvt_name ORLEFT_eqvt_name ih1 ih2 ih3 ih4 g)
       
  3325   }
       
  3326 next
       
  3327   case (IMP A B)
       
  3328   have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
       
  3329   have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
       
  3330   have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3331   have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3332   have g: "pi\<bullet>(\<parallel>(A IMP B)\<parallel>) = (\<parallel>(A IMP B)\<parallel>)"
       
  3333     apply -
       
  3334     apply(simp only: lfp_eqvt)
       
  3335     apply(simp only: comp_def)
       
  3336     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3337     apply(simp only: NEGc.simps NEGn.simps)
       
  3338     apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name IMPRIGHT_eqvt_name IMPLEFT_eqvt_name)
       
  3339     apply(perm_simp add: ih1 ih2 ih3 ih4)
       
  3340     done
       
  3341   { case 1 show ?case by (rule g)
       
  3342   next 
       
  3343     case 2 show ?case
       
  3344       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name 
       
  3345                      IMPRIGHT_eqvt_name IMPLEFT_eqvt_name ih1 ih2 ih3 ih4 g)
       
  3346   }
       
  3347 qed
       
  3348 
       
  3349 lemma CAND_eqvt_coname:
       
  3350   fixes pi::"coname prm"
       
  3351   shows   "(pi\<bullet>(\<parallel>(B)\<parallel>)) = (\<parallel>(B)\<parallel>)"
       
  3352   and     "(pi\<bullet>(\<parallel><B>\<parallel>)) = (\<parallel><B>\<parallel>)"
       
  3353 proof (nominal_induct B rule: ty.strong_induct)
       
  3354   case (PR X)
       
  3355   { case 1 show ?case 
       
  3356       apply -
       
  3357       apply(simp add: lfp_eqvt)
       
  3358       apply(simp add: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3359       apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname)
       
  3360       apply(perm_simp)
       
  3361     done
       
  3362   next
       
  3363     case 2 show ?case
       
  3364       apply -
       
  3365       apply(simp only: NEGc_simps)
       
  3366       apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname)
       
  3367       apply(simp add: lfp_eqvt)
       
  3368       apply(simp add: comp_def)
       
  3369       apply(simp add: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3370       apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname)
       
  3371       apply(perm_simp)
       
  3372       done
       
  3373   }
       
  3374 next
       
  3375   case (NOT B)
       
  3376   have ih1: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3377   have ih2: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3378   have g: "pi\<bullet>(\<parallel>(NOT B)\<parallel>) = (\<parallel>(NOT B)\<parallel>)"
       
  3379     apply -
       
  3380     apply(simp only: lfp_eqvt)
       
  3381     apply(simp only: comp_def)
       
  3382     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3383     apply(simp only: NEGc.simps NEGn.simps)
       
  3384     apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
       
  3385             NOTRIGHT_eqvt_coname NOTLEFT_eqvt_coname)
       
  3386     apply(perm_simp add: ih1 ih2)
       
  3387     done
       
  3388   { case 1 show ?case by (rule g)
       
  3389   next 
       
  3390     case 2 show ?case
       
  3391       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
       
  3392               NOTRIGHT_eqvt_coname ih1 ih2 g)
       
  3393   }
       
  3394 next
       
  3395   case (AND A B)
       
  3396   have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
       
  3397   have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
       
  3398   have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3399   have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3400   have g: "pi\<bullet>(\<parallel>(A AND B)\<parallel>) = (\<parallel>(A AND B)\<parallel>)"
       
  3401     apply -
       
  3402     apply(simp only: lfp_eqvt)
       
  3403     apply(simp only: comp_def)
       
  3404     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3405     apply(simp only: NEGc.simps NEGn.simps)
       
  3406     apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname ANDRIGHT_eqvt_coname 
       
  3407                      ANDLEFT2_eqvt_coname ANDLEFT1_eqvt_coname)
       
  3408     apply(perm_simp add: ih1 ih2 ih3 ih4)
       
  3409     done
       
  3410   { case 1 show ?case by (rule g)
       
  3411   next 
       
  3412     case 2 show ?case
       
  3413       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
       
  3414                      ANDRIGHT_eqvt_coname ANDLEFT1_eqvt_coname ANDLEFT2_eqvt_coname ih1 ih2 ih3 ih4 g)
       
  3415   }
       
  3416 next
       
  3417   case (OR A B)
       
  3418   have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
       
  3419   have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
       
  3420   have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3421   have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3422   have g: "pi\<bullet>(\<parallel>(A OR B)\<parallel>) = (\<parallel>(A OR B)\<parallel>)"
       
  3423     apply -
       
  3424     apply(simp only: lfp_eqvt)
       
  3425     apply(simp only: comp_def)
       
  3426     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3427     apply(simp only: NEGc.simps NEGn.simps)
       
  3428     apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname ORRIGHT1_eqvt_coname 
       
  3429                      ORRIGHT2_eqvt_coname ORLEFT_eqvt_coname)
       
  3430     apply(perm_simp add: ih1 ih2 ih3 ih4)
       
  3431     done
       
  3432   { case 1 show ?case by (rule g)
       
  3433   next 
       
  3434     case 2 show ?case
       
  3435       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
       
  3436                      ORRIGHT1_eqvt_coname ORRIGHT2_eqvt_coname ORLEFT_eqvt_coname ih1 ih2 ih3 ih4 g)
       
  3437   }
       
  3438 next
       
  3439   case (IMP A B)
       
  3440   have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
       
  3441   have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
       
  3442   have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
       
  3443   have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
       
  3444   have g: "pi\<bullet>(\<parallel>(A IMP B)\<parallel>) = (\<parallel>(A IMP B)\<parallel>)"
       
  3445     apply -
       
  3446     apply(simp only: lfp_eqvt)
       
  3447     apply(simp only: comp_def)
       
  3448     apply(simp only: perm_fun_def [where 'a="ntrm \<Rightarrow> bool"])
       
  3449     apply(simp only: NEGc.simps NEGn.simps)
       
  3450     apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname IMPRIGHT_eqvt_coname 
       
  3451          IMPLEFT_eqvt_coname)
       
  3452     apply(perm_simp add: ih1 ih2 ih3 ih4)
       
  3453     done
       
  3454   { case 1 show ?case by (rule g)
       
  3455   next 
       
  3456     case 2 show ?case
       
  3457       by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
       
  3458                      IMPRIGHT_eqvt_coname IMPLEFT_eqvt_coname ih1 ih2 ih3 ih4 g)
       
  3459   }
       
  3460 qed
       
  3461 
       
  3462 text {* Elimination rules for the set-operators *}
       
  3463 
       
  3464 lemma BINDINGc_elim:
       
  3465   assumes a: "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  3466   shows "\<forall>x P. ((x):P)\<in>(\<parallel>(B)\<parallel>) \<longrightarrow> SNa (M{a:=(x).P})"
       
  3467 using a
       
  3468 apply(auto simp add: BINDINGc_def)
       
  3469 apply(auto simp add: ctrm.inject alpha)
       
  3470 apply(drule_tac x="[(a,aa)]\<bullet>x" in spec)
       
  3471 apply(drule_tac x="[(a,aa)]\<bullet>P" in spec)
       
  3472 apply(drule mp)
       
  3473 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3474 apply(simp add: CAND_eqvt_coname)
       
  3475 apply(drule_tac ?pi2.0="[(a,aa)]" in SNa_eqvt(2))
       
  3476 apply(perm_simp add: csubst_eqvt)
       
  3477 done
       
  3478 
       
  3479 lemma BINDINGn_elim:
       
  3480   assumes a: "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
       
  3481   shows "\<forall>c P. (<c>:P)\<in>(\<parallel><B>\<parallel>) \<longrightarrow> SNa (M{x:=<c>.P})"
       
  3482 using a
       
  3483 apply(auto simp add: BINDINGn_def)
       
  3484 apply(auto simp add: ntrm.inject alpha)
       
  3485 apply(drule_tac x="[(x,xa)]\<bullet>c" in spec)
       
  3486 apply(drule_tac x="[(x,xa)]\<bullet>P" in spec)
       
  3487 apply(drule mp)
       
  3488 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3489 apply(simp add: CAND_eqvt_name)
       
  3490 apply(drule_tac ?pi1.0="[(x,xa)]" in SNa_eqvt(1))
       
  3491 apply(perm_simp add: nsubst_eqvt)
       
  3492 done
       
  3493 
       
  3494 lemma NOTRIGHT_elim:
       
  3495   assumes a: "<a>:M \<in> NOTRIGHT (NOT B) (\<parallel>(B)\<parallel>)"
       
  3496   obtains x' M' where "M = NotR (x').M' a" and "fic (NotR (x').M' a) a" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
       
  3497 using a
       
  3498 apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
       
  3499 apply(drule_tac x="x" in meta_spec)
       
  3500 apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
       
  3501 apply(simp)
       
  3502 apply(drule meta_mp)
       
  3503 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
       
  3504 apply(simp add: calc_atm)
       
  3505 apply(drule meta_mp)
       
  3506 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3507 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3508 apply(simp)
       
  3509 done
       
  3510 
       
  3511 lemma NOTLEFT_elim:
       
  3512   assumes a: "(x):M \<in> NOTLEFT (NOT B) (\<parallel><B>\<parallel>)"
       
  3513   obtains a' M' where "M = NotL <a'>.M' x" and "fin (NotL <a'>.M' x) x" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)"
       
  3514 using a
       
  3515 apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
       
  3516 apply(drule_tac x="a" in meta_spec)
       
  3517 apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
       
  3518 apply(simp)
       
  3519 apply(drule meta_mp)
       
  3520 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
       
  3521 apply(simp add: calc_atm)
       
  3522 apply(drule meta_mp)
       
  3523 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3524 apply(simp add: calc_atm CAND_eqvt_name)
       
  3525 apply(simp)
       
  3526 done
       
  3527 
       
  3528 lemma ANDRIGHT_elim:
       
  3529   assumes a: "<a>:M \<in> ANDRIGHT (B AND C) (\<parallel><B>\<parallel>) (\<parallel><C>\<parallel>)"
       
  3530   obtains d' M' e' N' where "M = AndR <d'>.M' <e'>.N' a" and "fic (AndR <d'>.M' <e'>.N' a) a" 
       
  3531                       and "<d'>:M' \<in> (\<parallel><B>\<parallel>)" and "<e'>:N' \<in> (\<parallel><C>\<parallel>)"
       
  3532 using a
       
  3533 apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm fresh_atm)
       
  3534 apply(drule_tac x="c" in meta_spec)
       
  3535 apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
       
  3536 apply(drule_tac x="c" in meta_spec)
       
  3537 apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
       
  3538 apply(simp)
       
  3539 apply(drule meta_mp)
       
  3540 apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
       
  3541 apply(simp add: calc_atm)
       
  3542 apply(drule meta_mp)
       
  3543 apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3544 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3545 apply(drule meta_mp)
       
  3546 apply(drule_tac pi="[(a,c)]" and x="<a>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3547 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3548 apply(simp)
       
  3549 apply(case_tac "a=b")
       
  3550 apply(simp)
       
  3551 apply(drule_tac x="c" in meta_spec)
       
  3552 apply(drule_tac x="[(b,c)]\<bullet>M" in meta_spec)
       
  3553 apply(drule_tac x="c" in meta_spec)
       
  3554 apply(drule_tac x="[(b,c)]\<bullet>N" in meta_spec)
       
  3555 apply(simp)
       
  3556 apply(drule meta_mp)
       
  3557 apply(drule_tac pi="[(b,c)]" in fic.eqvt(2))
       
  3558 apply(simp add: calc_atm)
       
  3559 apply(drule meta_mp)
       
  3560 apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3561 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3562 apply(drule meta_mp)
       
  3563 apply(drule_tac pi="[(b,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3564 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3565 apply(simp)
       
  3566 apply(simp)
       
  3567 apply(case_tac "c=b")
       
  3568 apply(simp)
       
  3569 apply(drule_tac x="b" in meta_spec)
       
  3570 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  3571 apply(drule_tac x="a" in meta_spec)
       
  3572 apply(drule_tac x="[(a,b)]\<bullet>N" in meta_spec)
       
  3573 apply(simp)
       
  3574 apply(drule meta_mp)
       
  3575 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  3576 apply(simp add: calc_atm)
       
  3577 apply(drule meta_mp)
       
  3578 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3579 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3580 apply(drule meta_mp)
       
  3581 apply(drule_tac pi="[(a,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3582 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3583 apply(simp)
       
  3584 apply(simp)
       
  3585 apply(drule_tac x="c" in meta_spec)
       
  3586 apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
       
  3587 apply(drule_tac x="b" in meta_spec)
       
  3588 apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
       
  3589 apply(simp)
       
  3590 apply(drule meta_mp)
       
  3591 apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
       
  3592 apply(simp add: calc_atm)
       
  3593 apply(drule meta_mp)
       
  3594 apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3595 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3596 apply(drule meta_mp)
       
  3597 apply(drule_tac pi="[(a,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3598 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3599 apply(simp)
       
  3600 apply(case_tac "a=aa")
       
  3601 apply(simp)
       
  3602 apply(drule_tac x="c" in meta_spec)
       
  3603 apply(drule_tac x="[(aa,c)]\<bullet>M" in meta_spec)
       
  3604 apply(drule_tac x="c" in meta_spec)
       
  3605 apply(drule_tac x="[(aa,c)]\<bullet>N" in meta_spec)
       
  3606 apply(simp)
       
  3607 apply(drule meta_mp)
       
  3608 apply(drule_tac pi="[(aa,c)]" in fic.eqvt(2))
       
  3609 apply(simp add: calc_atm)
       
  3610 apply(drule meta_mp)
       
  3611 apply(drule_tac pi="[(aa,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3612 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3613 apply(drule meta_mp)
       
  3614 apply(drule_tac pi="[(aa,c)]" and x="<aa>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3615 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3616 apply(simp)
       
  3617 apply(simp)
       
  3618 apply(case_tac "c=aa")
       
  3619 apply(simp)
       
  3620 apply(drule_tac x="a" in meta_spec)
       
  3621 apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
       
  3622 apply(drule_tac x="aa" in meta_spec)
       
  3623 apply(drule_tac x="[(a,aa)]\<bullet>N" in meta_spec)
       
  3624 apply(simp)
       
  3625 apply(drule meta_mp)
       
  3626 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
       
  3627 apply(simp add: calc_atm)
       
  3628 apply(drule meta_mp)
       
  3629 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3630 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3631 apply(drule meta_mp)
       
  3632 apply(drule_tac pi="[(a,aa)]" and x="<a>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3633 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3634 apply(simp)
       
  3635 apply(simp)
       
  3636 apply(drule_tac x="aa" in meta_spec)
       
  3637 apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
       
  3638 apply(drule_tac x="c" in meta_spec)
       
  3639 apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
       
  3640 apply(simp)
       
  3641 apply(drule meta_mp)
       
  3642 apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
       
  3643 apply(simp add: calc_atm)
       
  3644 apply(drule meta_mp)
       
  3645 apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3646 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3647 apply(drule meta_mp)
       
  3648 apply(drule_tac pi="[(a,c)]" and x="<a>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3649 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3650 apply(simp)
       
  3651 apply(case_tac "a=aa")
       
  3652 apply(simp)
       
  3653 apply(case_tac "aa=b")
       
  3654 apply(simp)
       
  3655 apply(drule_tac x="c" in meta_spec)
       
  3656 apply(drule_tac x="[(b,c)]\<bullet>M" in meta_spec)
       
  3657 apply(drule_tac x="c" in meta_spec)
       
  3658 apply(drule_tac x="[(b,c)]\<bullet>N" in meta_spec)
       
  3659 apply(simp)
       
  3660 apply(drule meta_mp)
       
  3661 apply(drule_tac pi="[(b,c)]" in fic.eqvt(2))
       
  3662 apply(simp add: calc_atm)
       
  3663 apply(drule meta_mp)
       
  3664 apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3665 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3666 apply(drule meta_mp)
       
  3667 apply(drule_tac pi="[(b,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3668 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3669 apply(simp)
       
  3670 apply(simp)
       
  3671 apply(case_tac "c=b")
       
  3672 apply(simp)
       
  3673 apply(drule_tac x="b" in meta_spec)
       
  3674 apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
       
  3675 apply(drule_tac x="aa" in meta_spec)
       
  3676 apply(drule_tac x="[(aa,b)]\<bullet>N" in meta_spec)
       
  3677 apply(simp)
       
  3678 apply(drule meta_mp)
       
  3679 apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
       
  3680 apply(simp add: calc_atm)
       
  3681 apply(drule meta_mp)
       
  3682 apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3683 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3684 apply(drule meta_mp)
       
  3685 apply(drule_tac pi="[(aa,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3686 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3687 apply(simp)
       
  3688 apply(simp)
       
  3689 apply(drule_tac x="c" in meta_spec)
       
  3690 apply(drule_tac x="[(aa,c)]\<bullet>M" in meta_spec)
       
  3691 apply(drule_tac x="b" in meta_spec)
       
  3692 apply(drule_tac x="[(aa,c)]\<bullet>N" in meta_spec)
       
  3693 apply(simp)
       
  3694 apply(drule meta_mp)
       
  3695 apply(drule_tac pi="[(aa,c)]" in fic.eqvt(2))
       
  3696 apply(simp add: calc_atm)
       
  3697 apply(drule meta_mp)
       
  3698 apply(drule_tac pi="[(aa,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3699 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3700 apply(drule meta_mp)
       
  3701 apply(drule_tac pi="[(aa,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3702 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3703 apply(simp)
       
  3704 apply(simp)
       
  3705 apply(case_tac "c=aa")
       
  3706 apply(simp)
       
  3707 apply(case_tac "a=b")
       
  3708 apply(simp)
       
  3709 apply(drule_tac x="b" in meta_spec)
       
  3710 apply(drule_tac x="[(b,aa)]\<bullet>M" in meta_spec)
       
  3711 apply(drule_tac x="aa" in meta_spec)
       
  3712 apply(drule_tac x="[(b,aa)]\<bullet>N" in meta_spec)
       
  3713 apply(simp)
       
  3714 apply(drule meta_mp)
       
  3715 apply(drule_tac pi="[(b,aa)]" in fic.eqvt(2))
       
  3716 apply(simp add: calc_atm)
       
  3717 apply(drule meta_mp)
       
  3718 apply(drule_tac pi="[(b,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3719 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3720 apply(drule meta_mp)
       
  3721 apply(drule_tac pi="[(b,aa)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3722 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3723 apply(simp)
       
  3724 apply(simp)
       
  3725 apply(case_tac "aa=b")
       
  3726 apply(simp)
       
  3727 apply(drule_tac x="a" in meta_spec)
       
  3728 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  3729 apply(drule_tac x="a" in meta_spec)
       
  3730 apply(drule_tac x="[(a,b)]\<bullet>N" in meta_spec)
       
  3731 apply(simp)
       
  3732 apply(drule meta_mp)
       
  3733 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  3734 apply(simp add: calc_atm)
       
  3735 apply(drule meta_mp)
       
  3736 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3737 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3738 apply(drule meta_mp)
       
  3739 apply(drule_tac pi="[(a,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3740 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3741 apply(simp)
       
  3742 apply(simp)
       
  3743 apply(drule_tac x="a" in meta_spec)
       
  3744 apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
       
  3745 apply(drule_tac x="b" in meta_spec)
       
  3746 apply(drule_tac x="[(a,aa)]\<bullet>N" in meta_spec)
       
  3747 apply(simp)
       
  3748 apply(drule meta_mp)
       
  3749 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
       
  3750 apply(simp add: calc_atm)
       
  3751 apply(drule meta_mp)
       
  3752 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3753 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3754 apply(drule meta_mp)
       
  3755 apply(drule_tac pi="[(a,aa)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3756 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3757 apply(simp)
       
  3758 apply(simp)
       
  3759 apply(case_tac "a=b")
       
  3760 apply(simp)
       
  3761 apply(drule_tac x="aa" in meta_spec)
       
  3762 apply(drule_tac x="[(b,c)]\<bullet>M" in meta_spec)
       
  3763 apply(drule_tac x="c" in meta_spec)
       
  3764 apply(drule_tac x="[(b,c)]\<bullet>N" in meta_spec)
       
  3765 apply(simp)
       
  3766 apply(drule meta_mp)
       
  3767 apply(drule_tac pi="[(b,c)]" in fic.eqvt(2))
       
  3768 apply(simp add: calc_atm)
       
  3769 apply(drule meta_mp)
       
  3770 apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3771 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3772 apply(drule meta_mp)
       
  3773 apply(drule_tac pi="[(b,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3774 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3775 apply(simp)
       
  3776 apply(simp)
       
  3777 apply(case_tac "c=b")
       
  3778 apply(simp)
       
  3779 apply(drule_tac x="aa" in meta_spec)
       
  3780 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  3781 apply(drule_tac x="a" in meta_spec)
       
  3782 apply(drule_tac x="[(a,b)]\<bullet>N" in meta_spec)
       
  3783 apply(simp)
       
  3784 apply(drule meta_mp)
       
  3785 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  3786 apply(simp add: calc_atm)
       
  3787 apply(drule meta_mp)
       
  3788 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3789 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3790 apply(drule meta_mp)
       
  3791 apply(drule_tac pi="[(a,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3792 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3793 apply(simp)
       
  3794 apply(simp)
       
  3795 apply(drule_tac x="aa" in meta_spec)
       
  3796 apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
       
  3797 apply(drule_tac x="b" in meta_spec)
       
  3798 apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
       
  3799 apply(simp)
       
  3800 apply(drule meta_mp)
       
  3801 apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
       
  3802 apply(simp add: calc_atm)
       
  3803 apply(drule meta_mp)
       
  3804 apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3805 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3806 apply(drule meta_mp)
       
  3807 apply(drule_tac pi="[(a,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3808 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3809 apply(simp)
       
  3810 done 
       
  3811 
       
  3812 lemma ANDLEFT1_elim:
       
  3813   assumes a: "(x):M \<in> ANDLEFT1 (B AND C) (\<parallel>(B)\<parallel>)"
       
  3814   obtains x' M' where "M = AndL1 (x').M' x" and "fin (AndL1 (x').M' x) x" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
       
  3815 using a
       
  3816 apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
       
  3817 apply(drule_tac x="y" in meta_spec)
       
  3818 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  3819 apply(simp)
       
  3820 apply(drule meta_mp)
       
  3821 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  3822 apply(simp add: calc_atm)
       
  3823 apply(drule meta_mp)
       
  3824 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3825 apply(simp add: calc_atm CAND_eqvt_name)
       
  3826 apply(simp)
       
  3827 apply(case_tac "x=xa")
       
  3828 apply(simp)
       
  3829 apply(drule_tac x="y" in meta_spec)
       
  3830 apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
       
  3831 apply(simp)
       
  3832 apply(drule meta_mp)
       
  3833 apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
       
  3834 apply(simp add: calc_atm)
       
  3835 apply(drule meta_mp)
       
  3836 apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3837 apply(simp add: calc_atm CAND_eqvt_name)
       
  3838 apply(simp)
       
  3839 apply(simp)
       
  3840 apply(case_tac "y=xa")
       
  3841 apply(simp)
       
  3842 apply(drule_tac x="x" in meta_spec)
       
  3843 apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
       
  3844 apply(simp)
       
  3845 apply(drule meta_mp)
       
  3846 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
       
  3847 apply(simp add: calc_atm)
       
  3848 apply(drule meta_mp)
       
  3849 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3850 apply(simp add: calc_atm CAND_eqvt_name)
       
  3851 apply(simp)
       
  3852 apply(simp)
       
  3853 apply(drule_tac x="xa" in meta_spec)
       
  3854 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  3855 apply(simp)
       
  3856 apply(drule meta_mp)
       
  3857 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  3858 apply(simp add: calc_atm)
       
  3859 apply(drule meta_mp)
       
  3860 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3861 apply(simp add: calc_atm CAND_eqvt_name)
       
  3862 apply(simp)
       
  3863 done
       
  3864 
       
  3865 lemma ANDLEFT2_elim:
       
  3866   assumes a: "(x):M \<in> ANDLEFT2 (B AND C) (\<parallel>(C)\<parallel>)"
       
  3867   obtains x' M' where "M = AndL2 (x').M' x" and "fin (AndL2 (x').M' x) x" and "(x'):M' \<in> (\<parallel>(C)\<parallel>)"
       
  3868 using a
       
  3869 apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
       
  3870 apply(drule_tac x="y" in meta_spec)
       
  3871 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  3872 apply(simp)
       
  3873 apply(drule meta_mp)
       
  3874 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  3875 apply(simp add: calc_atm)
       
  3876 apply(drule meta_mp)
       
  3877 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3878 apply(simp add: calc_atm CAND_eqvt_name)
       
  3879 apply(simp)
       
  3880 apply(case_tac "x=xa")
       
  3881 apply(simp)
       
  3882 apply(drule_tac x="y" in meta_spec)
       
  3883 apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
       
  3884 apply(simp)
       
  3885 apply(drule meta_mp)
       
  3886 apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
       
  3887 apply(simp add: calc_atm)
       
  3888 apply(drule meta_mp)
       
  3889 apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3890 apply(simp add: calc_atm CAND_eqvt_name)
       
  3891 apply(simp)
       
  3892 apply(simp)
       
  3893 apply(case_tac "y=xa")
       
  3894 apply(simp)
       
  3895 apply(drule_tac x="x" in meta_spec)
       
  3896 apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
       
  3897 apply(simp)
       
  3898 apply(drule meta_mp)
       
  3899 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
       
  3900 apply(simp add: calc_atm)
       
  3901 apply(drule meta_mp)
       
  3902 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3903 apply(simp add: calc_atm CAND_eqvt_name)
       
  3904 apply(simp)
       
  3905 apply(simp)
       
  3906 apply(drule_tac x="xa" in meta_spec)
       
  3907 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  3908 apply(simp)
       
  3909 apply(drule meta_mp)
       
  3910 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  3911 apply(simp add: calc_atm)
       
  3912 apply(drule meta_mp)
       
  3913 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  3914 apply(simp add: calc_atm CAND_eqvt_name)
       
  3915 apply(simp)
       
  3916 done
       
  3917 
       
  3918 lemma ORRIGHT1_elim:
       
  3919   assumes a: "<a>:M \<in> ORRIGHT1 (B OR C) (\<parallel><B>\<parallel>)"
       
  3920   obtains a' M' where "M = OrR1 <a'>.M' a" and "fic (OrR1 <a'>.M' a) a" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)"
       
  3921 using a
       
  3922 apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
       
  3923 apply(drule_tac x="b" in meta_spec)
       
  3924 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  3925 apply(simp)
       
  3926 apply(drule meta_mp)
       
  3927 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  3928 apply(simp add: calc_atm)
       
  3929 apply(drule meta_mp)
       
  3930 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3931 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3932 apply(simp)
       
  3933 apply(case_tac "a=aa")
       
  3934 apply(simp)
       
  3935 apply(drule_tac x="b" in meta_spec)
       
  3936 apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
       
  3937 apply(simp)
       
  3938 apply(drule meta_mp)
       
  3939 apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
       
  3940 apply(simp add: calc_atm)
       
  3941 apply(drule meta_mp)
       
  3942 apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3943 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3944 apply(simp)
       
  3945 apply(simp)
       
  3946 apply(case_tac "b=aa")
       
  3947 apply(simp)
       
  3948 apply(drule_tac x="a" in meta_spec)
       
  3949 apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
       
  3950 apply(simp)
       
  3951 apply(drule meta_mp)
       
  3952 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
       
  3953 apply(simp add: calc_atm)
       
  3954 apply(drule meta_mp)
       
  3955 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3956 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3957 apply(simp)
       
  3958 apply(simp)
       
  3959 apply(drule_tac x="aa" in meta_spec)
       
  3960 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  3961 apply(simp)
       
  3962 apply(drule meta_mp)
       
  3963 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  3964 apply(simp add: calc_atm)
       
  3965 apply(drule meta_mp)
       
  3966 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3967 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3968 apply(simp)
       
  3969 done
       
  3970 
       
  3971 lemma ORRIGHT2_elim:
       
  3972   assumes a: "<a>:M \<in> ORRIGHT2 (B OR C) (\<parallel><C>\<parallel>)"
       
  3973   obtains a' M' where "M = OrR2 <a'>.M' a" and "fic (OrR2 <a'>.M' a) a" and "<a'>:M' \<in> (\<parallel><C>\<parallel>)"
       
  3974 using a
       
  3975 apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
       
  3976 apply(drule_tac x="b" in meta_spec)
       
  3977 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  3978 apply(simp)
       
  3979 apply(drule meta_mp)
       
  3980 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  3981 apply(simp add: calc_atm)
       
  3982 apply(drule meta_mp)
       
  3983 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3984 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3985 apply(simp)
       
  3986 apply(case_tac "a=aa")
       
  3987 apply(simp)
       
  3988 apply(drule_tac x="b" in meta_spec)
       
  3989 apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
       
  3990 apply(simp)
       
  3991 apply(drule meta_mp)
       
  3992 apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
       
  3993 apply(simp add: calc_atm)
       
  3994 apply(drule meta_mp)
       
  3995 apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  3996 apply(simp add: calc_atm CAND_eqvt_coname)
       
  3997 apply(simp)
       
  3998 apply(simp)
       
  3999 apply(case_tac "b=aa")
       
  4000 apply(simp)
       
  4001 apply(drule_tac x="a" in meta_spec)
       
  4002 apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
       
  4003 apply(simp)
       
  4004 apply(drule meta_mp)
       
  4005 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
       
  4006 apply(simp add: calc_atm)
       
  4007 apply(drule meta_mp)
       
  4008 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4009 apply(simp add: calc_atm CAND_eqvt_coname)
       
  4010 apply(simp)
       
  4011 apply(simp)
       
  4012 apply(drule_tac x="aa" in meta_spec)
       
  4013 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  4014 apply(simp)
       
  4015 apply(drule meta_mp)
       
  4016 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  4017 apply(simp add: calc_atm)
       
  4018 apply(drule meta_mp)
       
  4019 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4020 apply(simp add: calc_atm CAND_eqvt_coname)
       
  4021 apply(simp)
       
  4022 done
       
  4023 
       
  4024 lemma ORLEFT_elim:
       
  4025   assumes a: "(x):M \<in> ORLEFT (B OR C) (\<parallel>(B)\<parallel>) (\<parallel>(C)\<parallel>)"
       
  4026   obtains y' M' z' N' where "M = OrL (y').M' (z').N' x" and "fin (OrL (y').M' (z').N' x) x" 
       
  4027                       and "(y'):M' \<in> (\<parallel>(B)\<parallel>)" and "(z'):N' \<in> (\<parallel>(C)\<parallel>)"
       
  4028 using a
       
  4029 apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm fresh_atm)
       
  4030 apply(drule_tac x="z" in meta_spec)
       
  4031 apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
       
  4032 apply(drule_tac x="z" in meta_spec)
       
  4033 apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
       
  4034 apply(simp)
       
  4035 apply(drule meta_mp)
       
  4036 apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
       
  4037 apply(simp add: calc_atm)
       
  4038 apply(drule meta_mp)
       
  4039 apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4040 apply(simp add: calc_atm CAND_eqvt_name)
       
  4041 apply(drule meta_mp)
       
  4042 apply(drule_tac pi="[(x,z)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4043 apply(simp add: calc_atm CAND_eqvt_name)
       
  4044 apply(simp)
       
  4045 apply(case_tac "x=y")
       
  4046 apply(simp)
       
  4047 apply(drule_tac x="z" in meta_spec)
       
  4048 apply(drule_tac x="[(y,z)]\<bullet>M" in meta_spec)
       
  4049 apply(drule_tac x="z" in meta_spec)
       
  4050 apply(drule_tac x="[(y,z)]\<bullet>N" in meta_spec)
       
  4051 apply(simp)
       
  4052 apply(drule meta_mp)
       
  4053 apply(drule_tac pi="[(y,z)]" in fin.eqvt(1))
       
  4054 apply(simp add: calc_atm)
       
  4055 apply(drule meta_mp)
       
  4056 apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4057 apply(simp add: calc_atm CAND_eqvt_name)
       
  4058 apply(drule meta_mp)
       
  4059 apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4060 apply(simp add: calc_atm CAND_eqvt_name)
       
  4061 apply(simp)
       
  4062 apply(simp)
       
  4063 apply(case_tac "z=y")
       
  4064 apply(simp)
       
  4065 apply(drule_tac x="y" in meta_spec)
       
  4066 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  4067 apply(drule_tac x="x" in meta_spec)
       
  4068 apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
       
  4069 apply(simp)
       
  4070 apply(drule meta_mp)
       
  4071 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  4072 apply(simp add: calc_atm)
       
  4073 apply(drule meta_mp)
       
  4074 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4075 apply(simp add: calc_atm CAND_eqvt_name)
       
  4076 apply(drule meta_mp)
       
  4077 apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4078 apply(simp add: calc_atm CAND_eqvt_name)
       
  4079 apply(simp)
       
  4080 apply(simp)
       
  4081 apply(drule_tac x="z" in meta_spec)
       
  4082 apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
       
  4083 apply(drule_tac x="y" in meta_spec)
       
  4084 apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
       
  4085 apply(simp)
       
  4086 apply(drule meta_mp)
       
  4087 apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
       
  4088 apply(simp add: calc_atm)
       
  4089 apply(drule meta_mp)
       
  4090 apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4091 apply(simp add: calc_atm CAND_eqvt_name)
       
  4092 apply(drule meta_mp)
       
  4093 apply(drule_tac pi="[(x,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4094 apply(simp add: calc_atm CAND_eqvt_name)
       
  4095 apply(simp)
       
  4096 apply(case_tac "x=xa")
       
  4097 apply(simp)
       
  4098 apply(drule_tac x="z" in meta_spec)
       
  4099 apply(drule_tac x="[(xa,z)]\<bullet>M" in meta_spec)
       
  4100 apply(drule_tac x="z" in meta_spec)
       
  4101 apply(drule_tac x="[(xa,z)]\<bullet>N" in meta_spec)
       
  4102 apply(simp)
       
  4103 apply(drule meta_mp)
       
  4104 apply(drule_tac pi="[(xa,z)]" in fin.eqvt(1))
       
  4105 apply(simp add: calc_atm)
       
  4106 apply(drule meta_mp)
       
  4107 apply(drule_tac pi="[(xa,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4108 apply(simp add: calc_atm CAND_eqvt_name)
       
  4109 apply(drule meta_mp)
       
  4110 apply(drule_tac pi="[(xa,z)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4111 apply(simp add: calc_atm CAND_eqvt_name)
       
  4112 apply(simp)
       
  4113 apply(simp)
       
  4114 apply(case_tac "z=xa")
       
  4115 apply(simp)
       
  4116 apply(drule_tac x="x" in meta_spec)
       
  4117 apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
       
  4118 apply(drule_tac x="xa" in meta_spec)
       
  4119 apply(drule_tac x="[(x,xa)]\<bullet>N" in meta_spec)
       
  4120 apply(simp)
       
  4121 apply(drule meta_mp)
       
  4122 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
       
  4123 apply(simp add: calc_atm)
       
  4124 apply(drule meta_mp)
       
  4125 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4126 apply(simp add: calc_atm CAND_eqvt_name)
       
  4127 apply(drule meta_mp)
       
  4128 apply(drule_tac pi="[(x,xa)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4129 apply(simp add: calc_atm CAND_eqvt_name)
       
  4130 apply(simp)
       
  4131 apply(simp)
       
  4132 apply(drule_tac x="xa" in meta_spec)
       
  4133 apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
       
  4134 apply(drule_tac x="z" in meta_spec)
       
  4135 apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
       
  4136 apply(simp)
       
  4137 apply(drule meta_mp)
       
  4138 apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
       
  4139 apply(simp add: calc_atm)
       
  4140 apply(drule meta_mp)
       
  4141 apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4142 apply(simp add: calc_atm CAND_eqvt_name)
       
  4143 apply(drule meta_mp)
       
  4144 apply(drule_tac pi="[(x,z)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4145 apply(simp add: calc_atm CAND_eqvt_name)
       
  4146 apply(simp)
       
  4147 apply(case_tac "x=xa")
       
  4148 apply(simp)
       
  4149 apply(case_tac "xa=y")
       
  4150 apply(simp)
       
  4151 apply(drule_tac x="z" in meta_spec)
       
  4152 apply(drule_tac x="[(y,z)]\<bullet>M" in meta_spec)
       
  4153 apply(drule_tac x="z" in meta_spec)
       
  4154 apply(drule_tac x="[(y,z)]\<bullet>N" in meta_spec)
       
  4155 apply(simp)
       
  4156 apply(drule meta_mp)
       
  4157 apply(drule_tac pi="[(y,z)]" in fin.eqvt(1))
       
  4158 apply(simp add: calc_atm)
       
  4159 apply(drule meta_mp)
       
  4160 apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4161 apply(simp add: calc_atm CAND_eqvt_name)
       
  4162 apply(drule meta_mp)
       
  4163 apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4164 apply(simp add: calc_atm CAND_eqvt_name)
       
  4165 apply(simp)
       
  4166 apply(simp)
       
  4167 apply(case_tac "z=y")
       
  4168 apply(simp)
       
  4169 apply(drule_tac x="y" in meta_spec)
       
  4170 apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
       
  4171 apply(drule_tac x="xa" in meta_spec)
       
  4172 apply(drule_tac x="[(xa,y)]\<bullet>N" in meta_spec)
       
  4173 apply(simp)
       
  4174 apply(drule meta_mp)
       
  4175 apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
       
  4176 apply(simp add: calc_atm)
       
  4177 apply(drule meta_mp)
       
  4178 apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4179 apply(simp add: calc_atm CAND_eqvt_name)
       
  4180 apply(drule meta_mp)
       
  4181 apply(drule_tac pi="[(xa,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4182 apply(simp add: calc_atm CAND_eqvt_name)
       
  4183 apply(simp)
       
  4184 apply(simp)
       
  4185 apply(drule_tac x="z" in meta_spec)
       
  4186 apply(drule_tac x="[(xa,z)]\<bullet>M" in meta_spec)
       
  4187 apply(drule_tac x="y" in meta_spec)
       
  4188 apply(drule_tac x="[(xa,z)]\<bullet>N" in meta_spec)
       
  4189 apply(simp)
       
  4190 apply(drule meta_mp)
       
  4191 apply(drule_tac pi="[(xa,z)]" in fin.eqvt(1))
       
  4192 apply(simp add: calc_atm)
       
  4193 apply(drule meta_mp)
       
  4194 apply(drule_tac pi="[(xa,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4195 apply(simp add: calc_atm CAND_eqvt_name)
       
  4196 apply(drule meta_mp)
       
  4197 apply(drule_tac pi="[(xa,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4198 apply(simp add: calc_atm CAND_eqvt_name)
       
  4199 apply(simp)
       
  4200 apply(simp)
       
  4201 apply(case_tac "z=xa")
       
  4202 apply(simp)
       
  4203 apply(case_tac "x=y")
       
  4204 apply(simp)
       
  4205 apply(drule_tac x="y" in meta_spec)
       
  4206 apply(drule_tac x="[(y,xa)]\<bullet>M" in meta_spec)
       
  4207 apply(drule_tac x="xa" in meta_spec)
       
  4208 apply(drule_tac x="[(y,xa)]\<bullet>N" in meta_spec)
       
  4209 apply(simp)
       
  4210 apply(drule meta_mp)
       
  4211 apply(drule_tac pi="[(y,xa)]" in fin.eqvt(1))
       
  4212 apply(simp add: calc_atm)
       
  4213 apply(drule meta_mp)
       
  4214 apply(drule_tac pi="[(y,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4215 apply(simp add: calc_atm CAND_eqvt_name)
       
  4216 apply(drule meta_mp)
       
  4217 apply(drule_tac pi="[(y,xa)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4218 apply(simp add: calc_atm CAND_eqvt_name)
       
  4219 apply(simp)
       
  4220 apply(simp)
       
  4221 apply(case_tac "xa=y")
       
  4222 apply(simp)
       
  4223 apply(drule_tac x="x" in meta_spec)
       
  4224 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  4225 apply(drule_tac x="x" in meta_spec)
       
  4226 apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
       
  4227 apply(simp)
       
  4228 apply(drule meta_mp)
       
  4229 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  4230 apply(simp add: calc_atm)
       
  4231 apply(drule meta_mp)
       
  4232 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4233 apply(simp add: calc_atm CAND_eqvt_name)
       
  4234 apply(drule meta_mp)
       
  4235 apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4236 apply(simp add: calc_atm CAND_eqvt_name)
       
  4237 apply(simp)
       
  4238 apply(simp)
       
  4239 apply(drule_tac x="x" in meta_spec)
       
  4240 apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
       
  4241 apply(drule_tac x="y" in meta_spec)
       
  4242 apply(drule_tac x="[(x,xa)]\<bullet>N" in meta_spec)
       
  4243 apply(simp)
       
  4244 apply(drule meta_mp)
       
  4245 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
       
  4246 apply(simp add: calc_atm)
       
  4247 apply(drule meta_mp)
       
  4248 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4249 apply(simp add: calc_atm CAND_eqvt_name)
       
  4250 apply(drule meta_mp)
       
  4251 apply(drule_tac pi="[(x,xa)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4252 apply(simp add: calc_atm CAND_eqvt_name)
       
  4253 apply(simp)
       
  4254 apply(simp)
       
  4255 apply(case_tac "x=y")
       
  4256 apply(simp)
       
  4257 apply(drule_tac x="xa" in meta_spec)
       
  4258 apply(drule_tac x="[(y,z)]\<bullet>M" in meta_spec)
       
  4259 apply(drule_tac x="z" in meta_spec)
       
  4260 apply(drule_tac x="[(y,z)]\<bullet>N" in meta_spec)
       
  4261 apply(simp)
       
  4262 apply(drule meta_mp)
       
  4263 apply(drule_tac pi="[(y,z)]" in fin.eqvt(1))
       
  4264 apply(simp add: calc_atm)
       
  4265 apply(drule meta_mp)
       
  4266 apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4267 apply(simp add: calc_atm CAND_eqvt_name)
       
  4268 apply(drule meta_mp)
       
  4269 apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4270 apply(simp add: calc_atm CAND_eqvt_name)
       
  4271 apply(simp)
       
  4272 apply(simp)
       
  4273 apply(case_tac "z=y")
       
  4274 apply(simp)
       
  4275 apply(drule_tac x="xa" in meta_spec)
       
  4276 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  4277 apply(drule_tac x="x" in meta_spec)
       
  4278 apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
       
  4279 apply(simp)
       
  4280 apply(drule meta_mp)
       
  4281 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  4282 apply(simp add: calc_atm)
       
  4283 apply(drule meta_mp)
       
  4284 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4285 apply(simp add: calc_atm CAND_eqvt_name)
       
  4286 apply(drule meta_mp)
       
  4287 apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4288 apply(simp add: calc_atm CAND_eqvt_name)
       
  4289 apply(simp)
       
  4290 apply(simp)
       
  4291 apply(drule_tac x="xa" in meta_spec)
       
  4292 apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
       
  4293 apply(drule_tac x="y" in meta_spec)
       
  4294 apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
       
  4295 apply(simp)
       
  4296 apply(drule meta_mp)
       
  4297 apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
       
  4298 apply(simp add: calc_atm)
       
  4299 apply(drule meta_mp)
       
  4300 apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4301 apply(simp add: calc_atm CAND_eqvt_name)
       
  4302 apply(drule meta_mp)
       
  4303 apply(drule_tac pi="[(x,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4304 apply(simp add: calc_atm CAND_eqvt_name)
       
  4305 apply(simp)
       
  4306 done
       
  4307 
       
  4308 lemma IMPRIGHT_elim:
       
  4309   assumes a: "<a>:M \<in> IMPRIGHT (B IMP C) (\<parallel>(B)\<parallel>) (\<parallel><C>\<parallel>) (\<parallel>(C)\<parallel>) (\<parallel><B>\<parallel>)"
       
  4310   obtains x' a' M' where "M = ImpR (x').<a'>.M' a" and "fic (ImpR (x').<a'>.M' a) a" 
       
  4311                    and "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(C)\<parallel> \<longrightarrow> (x'):(M'{a':=(z).P}) \<in> \<parallel>(B)\<parallel>" 
       
  4312                    and "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><B>\<parallel> \<longrightarrow> <a'>:(M'{x':=<c>.Q}) \<in> \<parallel><C>\<parallel>"
       
  4313 using a
       
  4314 apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
       
  4315 apply(drule_tac x="x" in meta_spec)
       
  4316 apply(drule_tac x="b" in meta_spec)
       
  4317 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  4318 apply(simp)
       
  4319 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  4320 apply(simp add: calc_atm)
       
  4321 apply(drule meta_mp)
       
  4322 apply(auto)[1]
       
  4323 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4324 apply(simp add: calc_atm CAND_eqvt_coname)
       
  4325 apply(drule_tac x="z" in spec)
       
  4326 apply(drule_tac x="[(a,b)]\<bullet>P" in spec)
       
  4327 apply(simp add: fresh_prod fresh_left calc_atm)
       
  4328 apply(drule_tac pi="[(a,b)]" and x="(x):M{a:=(z).([(a,b)]\<bullet>P)}" 
       
  4329                                      in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4330 apply(perm_simp add: calc_atm csubst_eqvt CAND_eqvt_coname)
       
  4331 apply(drule meta_mp)
       
  4332 apply(auto)[1]
       
  4333 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4334 apply(simp add:  CAND_eqvt_coname)
       
  4335 apply(rotate_tac 2)
       
  4336 apply(drule_tac x="[(a,b)]\<bullet>c" in spec)
       
  4337 apply(drule_tac x="[(a,b)]\<bullet>Q" in spec)
       
  4338 apply(simp add: fresh_prod fresh_left)
       
  4339 apply(drule mp)
       
  4340 apply(simp add: calc_atm)
       
  4341 apply(drule_tac pi="[(a,b)]" and x="<a>:M{x:=<([(a,b)]\<bullet>c)>.([(a,b)]\<bullet>Q)}" 
       
  4342                                         in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4343 apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname)
       
  4344 apply(simp add: calc_atm)
       
  4345 apply(case_tac "a=aa")
       
  4346 apply(simp)
       
  4347 apply(drule_tac x="x" in meta_spec)
       
  4348 apply(drule_tac x="b" in meta_spec)
       
  4349 apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
       
  4350 apply(simp)
       
  4351 apply(drule meta_mp)
       
  4352 apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
       
  4353 apply(simp add: calc_atm)
       
  4354 apply(drule meta_mp)
       
  4355 apply(auto)[1]
       
  4356 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4357 apply(simp add: calc_atm CAND_eqvt_coname)
       
  4358 apply(drule_tac x="z" in spec)
       
  4359 apply(drule_tac x="[(a,b)]\<bullet>P" in spec)
       
  4360 apply(simp add: fresh_prod fresh_left calc_atm)
       
  4361 apply(drule_tac pi="[(a,b)]" and x="(x):M{a:=(z).([(a,b)]\<bullet>P)}" 
       
  4362                                      in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4363 apply(perm_simp add: calc_atm csubst_eqvt  CAND_eqvt_coname)
       
  4364 apply(drule meta_mp)
       
  4365 apply(auto)[1]
       
  4366 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4367 apply(simp add: CAND_eqvt_coname)
       
  4368 apply(drule_tac x="[(a,b)]\<bullet>c" in spec)
       
  4369 apply(drule_tac x="[(a,b)]\<bullet>Q" in spec)
       
  4370 apply(simp)
       
  4371 apply(simp add: fresh_prod fresh_left)
       
  4372 apply(drule mp)
       
  4373 apply(simp add: calc_atm)
       
  4374 apply(drule_tac pi="[(a,b)]" and x="<a>:M{x:=<([(a,b)]\<bullet>c)>.([(a,b)]\<bullet>Q)}" 
       
  4375                                       in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4376 apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname)
       
  4377 apply(simp add: calc_atm)
       
  4378 apply(simp)
       
  4379 apply(case_tac "b=aa")
       
  4380 apply(simp)
       
  4381 apply(drule_tac x="x" in meta_spec)
       
  4382 apply(drule_tac x="a" in meta_spec)
       
  4383 apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
       
  4384 apply(simp)
       
  4385 apply(drule meta_mp)
       
  4386 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
       
  4387 apply(simp add: calc_atm)
       
  4388 apply(drule meta_mp)
       
  4389 apply(auto)[1]
       
  4390 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4391 apply(simp add: calc_atm CAND_eqvt_coname)
       
  4392 apply(drule_tac x="z" in spec)
       
  4393 apply(drule_tac x="[(a,aa)]\<bullet>P" in spec)
       
  4394 apply(simp add: fresh_prod fresh_left calc_atm)
       
  4395 apply(drule_tac pi="[(a,aa)]" and x="(x):M{aa:=(z).([(a,aa)]\<bullet>P)}" 
       
  4396                                     in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4397 apply(perm_simp add: calc_atm csubst_eqvt  CAND_eqvt_coname)
       
  4398 apply(drule meta_mp)
       
  4399 apply(auto)[1]
       
  4400 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4401 apply(simp add:  CAND_eqvt_coname)
       
  4402 apply(drule_tac x="[(a,aa)]\<bullet>c" in spec)
       
  4403 apply(drule_tac x="[(a,aa)]\<bullet>Q" in spec)
       
  4404 apply(simp)
       
  4405 apply(simp add: fresh_prod fresh_left)
       
  4406 apply(drule mp)
       
  4407 apply(simp add: calc_atm)
       
  4408 apply(drule_tac pi="[(a,aa)]" and x="<aa>:M{x:=<([(a,aa)]\<bullet>c)>.([(a,aa)]\<bullet>Q)}" 
       
  4409                                     in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4410 apply(perm_simp add: nsubst_eqvt  CAND_eqvt_coname)
       
  4411 apply(simp add: calc_atm)
       
  4412 apply(simp)
       
  4413 apply(drule_tac x="x" in meta_spec)
       
  4414 apply(drule_tac x="aa" in meta_spec)
       
  4415 apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
       
  4416 apply(simp)
       
  4417 apply(drule meta_mp)
       
  4418 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  4419 apply(simp add: calc_atm)
       
  4420 apply(drule meta_mp)
       
  4421 apply(auto)[1]
       
  4422 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4423 apply(simp add: calc_atm  CAND_eqvt_coname)
       
  4424 apply(drule_tac x="z" in spec)
       
  4425 apply(drule_tac x="[(a,b)]\<bullet>P" in spec)
       
  4426 apply(simp add: fresh_prod fresh_left calc_atm)
       
  4427 apply(drule_tac pi="[(a,b)]" and x="(x):M{aa:=(z).([(a,b)]\<bullet>P)}" 
       
  4428                                           in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4429 apply(perm_simp add: calc_atm csubst_eqvt  CAND_eqvt_coname)
       
  4430 apply(drule meta_mp)
       
  4431 apply(auto)[1]
       
  4432 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4433 apply(simp add:  CAND_eqvt_coname)
       
  4434 apply(drule_tac x="[(a,b)]\<bullet>c" in spec)
       
  4435 apply(drule_tac x="[(a,b)]\<bullet>Q" in spec)
       
  4436 apply(simp add: fresh_prod fresh_left)
       
  4437 apply(drule mp)
       
  4438 apply(simp add: calc_atm)
       
  4439 apply(drule_tac pi="[(a,b)]" and x="<aa>:M{x:=<([(a,b)]\<bullet>c)>.([(a,b)]\<bullet>Q)}" 
       
  4440                                         in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4441 apply(perm_simp add: nsubst_eqvt  CAND_eqvt_coname)
       
  4442 apply(simp add: calc_atm)
       
  4443 done
       
  4444 
       
  4445 lemma IMPLEFT_elim:
       
  4446   assumes a: "(x):M \<in> IMPLEFT (B IMP C) (\<parallel><B>\<parallel>) (\<parallel>(C)\<parallel>)"
       
  4447   obtains x' a' M' N' where "M = ImpL <a'>.M' (x').N' x" and "fin (ImpL <a'>.M' (x').N' x) x" 
       
  4448                    and "<a'>:M' \<in> \<parallel><B>\<parallel>" and "(x'):N' \<in> \<parallel>(C)\<parallel>"
       
  4449 using a
       
  4450 apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
       
  4451 apply(drule_tac x="a" in meta_spec)
       
  4452 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  4453 apply(drule_tac x="y" in meta_spec)
       
  4454 apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
       
  4455 apply(simp)
       
  4456 apply(drule meta_mp)
       
  4457 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  4458 apply(simp add: calc_atm)
       
  4459 apply(drule meta_mp)
       
  4460 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4461 apply(simp add: calc_atm CAND_eqvt_name)
       
  4462 apply(drule meta_mp)
       
  4463 apply(drule_tac pi="[(x,y)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4464 apply(perm_simp add: calc_atm  CAND_eqvt_name)
       
  4465 apply(simp)
       
  4466 apply(case_tac "x=xa")
       
  4467 apply(simp)
       
  4468 apply(drule_tac x="a" in meta_spec)
       
  4469 apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
       
  4470 apply(drule_tac x="y" in meta_spec)
       
  4471 apply(drule_tac x="[(xa,y)]\<bullet>N" in meta_spec)
       
  4472 apply(simp)
       
  4473 apply(drule meta_mp)
       
  4474 apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
       
  4475 apply(simp add: calc_atm)
       
  4476 apply(drule meta_mp)
       
  4477 apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4478 apply(simp add: calc_atm CAND_eqvt_name)
       
  4479 apply(drule meta_mp)
       
  4480 apply(drule_tac pi="[(xa,y)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4481 apply(simp add: calc_atm CAND_eqvt_name)
       
  4482 apply(simp)
       
  4483 apply(simp)
       
  4484 apply(case_tac "y=xa")
       
  4485 apply(simp)
       
  4486 apply(drule_tac x="a" in meta_spec)
       
  4487 apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
       
  4488 apply(drule_tac x="x" in meta_spec)
       
  4489 apply(drule_tac x="[(x,xa)]\<bullet>N" in meta_spec)
       
  4490 apply(simp)
       
  4491 apply(drule meta_mp)
       
  4492 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
       
  4493 apply(simp add: calc_atm)
       
  4494 apply(drule meta_mp)
       
  4495 apply(drule_tac pi="[(x,xa)]" and x="<a>:M" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4496 apply(simp add: calc_atm  CAND_eqvt_name)
       
  4497 apply(drule meta_mp)
       
  4498 apply(drule_tac pi="[(x,xa)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4499 apply(simp add: calc_atm CAND_eqvt_name)
       
  4500 apply(simp)
       
  4501 apply(simp)
       
  4502 apply(drule_tac x="a" in meta_spec)
       
  4503 apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
       
  4504 apply(drule_tac x="xa" in meta_spec)
       
  4505 apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
       
  4506 apply(simp)
       
  4507 apply(drule meta_mp)
       
  4508 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  4509 apply(simp add: calc_atm)
       
  4510 apply(drule meta_mp)
       
  4511 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4512 apply(simp add: calc_atm CAND_eqvt_name)
       
  4513 apply(drule meta_mp)
       
  4514 apply(drule_tac pi="[(x,y)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4515 apply(simp add: calc_atm CAND_eqvt_name)
       
  4516 apply(simp)
       
  4517 done
       
  4518 
       
  4519 lemma CANDs_alpha:
       
  4520   shows "<a>:M \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> [a].M = [b].N \<Longrightarrow> <b>:N \<in> (\<parallel><B>\<parallel>)"
       
  4521   and   "(x):M \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> [x].M = [y].N \<Longrightarrow> (y):N \<in> (\<parallel>(B)\<parallel>)"
       
  4522 apply(auto simp add: alpha)
       
  4523 apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4524 apply(perm_simp add: CAND_eqvt_coname calc_atm)
       
  4525 apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4526 apply(perm_simp add: CAND_eqvt_name calc_atm)
       
  4527 done
       
  4528 
       
  4529 lemma CAND_NotR_elim:
       
  4530   assumes a: "<a>:NotR (x).M a \<in> (\<parallel><B>\<parallel>)" "<a>:NotR (x).M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  4531   shows "\<exists>B'. B = NOT B' \<and> (x):M \<in> (\<parallel>(B')\<parallel>)" 
       
  4532 using a
       
  4533 apply(nominal_induct B rule: ty.strong_induct)
       
  4534 apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
       
  4535 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4536 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4537 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4538 done
       
  4539 
       
  4540 lemma CAND_NotL_elim_aux:
       
  4541   assumes a: "(x):NotL <a>.M x \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):NotL <a>.M x \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
       
  4542   shows "\<exists>B'. B = NOT B' \<and> <a>:M \<in> (\<parallel><B'>\<parallel>)" 
       
  4543 using a
       
  4544 apply(nominal_induct B rule: ty.strong_induct)
       
  4545 apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
       
  4546 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4547 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4548 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4549 done
       
  4550 
       
  4551 lemmas CAND_NotL_elim = CAND_NotL_elim_aux[OF NEG_elim(2)]
       
  4552 
       
  4553 lemma CAND_AndR_elim:
       
  4554   assumes a: "<a>:AndR <b>.M <c>.N a \<in> (\<parallel><B>\<parallel>)" "<a>:AndR <b>.M <c>.N a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  4555   shows "\<exists>B1 B2. B = B1 AND B2 \<and> <b>:M \<in> (\<parallel><B1>\<parallel>) \<and> <c>:N \<in> (\<parallel><B2>\<parallel>)" 
       
  4556 using a
       
  4557 apply(nominal_induct B rule: ty.strong_induct)
       
  4558 apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
       
  4559 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4560 apply(drule_tac pi="[(a,ca)]" and x="<a>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4561 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4562 apply(auto intro: CANDs_alpha)[1]
       
  4563 apply(drule_tac pi="[(a,ca)]" and x="<a>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4564 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4565 apply(auto intro: CANDs_alpha)[1]
       
  4566 apply(drule_tac pi="[(a,ca)]" and x="<a>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4567 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4568 apply(auto intro: CANDs_alpha)[1]
       
  4569 apply(case_tac "a=ba")
       
  4570 apply(simp)
       
  4571 apply(drule_tac pi="[(ba,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4572 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4573 apply(auto intro: CANDs_alpha)[1]
       
  4574 apply(simp)
       
  4575 apply(case_tac "ca=ba")
       
  4576 apply(simp)
       
  4577 apply(drule_tac pi="[(a,ba)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4578 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4579 apply(auto intro: CANDs_alpha)[1]
       
  4580 apply(simp)
       
  4581 apply(drule_tac pi="[(a,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4582 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4583 apply(auto intro: CANDs_alpha)[1]
       
  4584 apply(case_tac "a=aa")
       
  4585 apply(simp)
       
  4586 apply(drule_tac pi="[(aa,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4587 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4588 apply(auto intro: CANDs_alpha)[1]
       
  4589 apply(simp)
       
  4590 apply(case_tac "ca=aa")
       
  4591 apply(simp)
       
  4592 apply(drule_tac pi="[(a,aa)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4593 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4594 apply(auto intro: CANDs_alpha)[1]
       
  4595 apply(simp)
       
  4596 apply(drule_tac pi="[(a,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4597 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4598 apply(auto intro: CANDs_alpha)[1]
       
  4599 apply(drule_tac pi="[(a,ca)]" and x="<a>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4600 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4601 apply(auto intro: CANDs_alpha)[1]
       
  4602 apply(case_tac "a=aa")
       
  4603 apply(simp)
       
  4604 apply(drule_tac pi="[(aa,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4605 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4606 apply(auto intro: CANDs_alpha)[1]
       
  4607 apply(simp)
       
  4608 apply(case_tac "ca=aa")
       
  4609 apply(simp)
       
  4610 apply(drule_tac pi="[(a,aa)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4611 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4612 apply(auto intro: CANDs_alpha)[1]
       
  4613 apply(simp)
       
  4614 apply(drule_tac pi="[(a,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4615 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4616 apply(auto intro: CANDs_alpha)[1]
       
  4617 apply(case_tac "a=ba")
       
  4618 apply(simp)
       
  4619 apply(drule_tac pi="[(ba,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4620 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4621 apply(auto intro: CANDs_alpha)[1]
       
  4622 apply(simp)
       
  4623 apply(case_tac "ca=ba")
       
  4624 apply(simp)
       
  4625 apply(drule_tac pi="[(a,ba)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4626 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4627 apply(auto intro: CANDs_alpha)[1]
       
  4628 apply(simp)
       
  4629 apply(drule_tac pi="[(a,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4630 apply(simp add: CAND_eqvt_coname calc_atm)
       
  4631 apply(auto intro: CANDs_alpha)[1]
       
  4632 done
       
  4633 
       
  4634 lemma CAND_OrR1_elim:
       
  4635   assumes a: "<a>:OrR1 <b>.M a \<in> (\<parallel><B>\<parallel>)" "<a>:OrR1 <b>.M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  4636   shows "\<exists>B1 B2. B = B1 OR B2 \<and> <b>:M \<in> (\<parallel><B1>\<parallel>)" 
       
  4637 using a
       
  4638 apply(nominal_induct B rule: ty.strong_induct)
       
  4639 apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
       
  4640 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4641 apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4642 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4643 apply(case_tac "a=aa")
       
  4644 apply(simp)
       
  4645 apply(drule_tac pi="[(aa,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4646 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4647 apply(case_tac "ba=aa")
       
  4648 apply(simp)
       
  4649 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4650 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4651 apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4652 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4653 done
       
  4654 
       
  4655 lemma CAND_OrR2_elim:
       
  4656   assumes a: "<a>:OrR2 <b>.M a \<in> (\<parallel><B>\<parallel>)" "<a>:OrR2 <b>.M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  4657   shows "\<exists>B1 B2. B = B1 OR B2 \<and> <b>:M \<in> (\<parallel><B2>\<parallel>)" 
       
  4658 using a
       
  4659 apply(nominal_induct B rule: ty.strong_induct)
       
  4660 apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
       
  4661 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4662 apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4663 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4664 apply(case_tac "a=aa")
       
  4665 apply(simp)
       
  4666 apply(drule_tac pi="[(aa,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4667 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4668 apply(case_tac "ba=aa")
       
  4669 apply(simp)
       
  4670 apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4671 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4672 apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4673 apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
       
  4674 done
       
  4675 
       
  4676 lemma CAND_OrL_elim_aux:
       
  4677   assumes a: "(x):(OrL (y).M (z).N x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(OrL (y).M (z).N x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
       
  4678   shows "\<exists>B1 B2. B = B1 OR B2 \<and> (y):M \<in> (\<parallel>(B1)\<parallel>) \<and> (z):N \<in> (\<parallel>(B2)\<parallel>)" 
       
  4679 using a
       
  4680 apply(nominal_induct B rule: ty.strong_induct)
       
  4681 apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
       
  4682 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4683 apply(drule_tac pi="[(x,za)]" and x="(x):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4684 apply(simp add: CAND_eqvt_name calc_atm)
       
  4685 apply(auto intro: CANDs_alpha)[1]
       
  4686 apply(drule_tac pi="[(x,za)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4687 apply(simp add: CAND_eqvt_name calc_atm)
       
  4688 apply(auto intro: CANDs_alpha)[1]
       
  4689 apply(drule_tac pi="[(x,za)]" and x="(x):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4690 apply(simp add: CAND_eqvt_name calc_atm)
       
  4691 apply(auto intro: CANDs_alpha)[1]
       
  4692 apply(case_tac "x=ya")
       
  4693 apply(simp)
       
  4694 apply(drule_tac pi="[(ya,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4695 apply(simp add: CAND_eqvt_name calc_atm)
       
  4696 apply(auto intro: CANDs_alpha)[1]
       
  4697 apply(simp)
       
  4698 apply(case_tac "za=ya")
       
  4699 apply(simp)
       
  4700 apply(drule_tac pi="[(x,ya)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4701 apply(simp add: CAND_eqvt_name calc_atm)
       
  4702 apply(auto intro: CANDs_alpha)[1]
       
  4703 apply(simp)
       
  4704 apply(drule_tac pi="[(x,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4705 apply(simp add: CAND_eqvt_name calc_atm)
       
  4706 apply(auto intro: CANDs_alpha)[1]
       
  4707 apply(case_tac "x=xa")
       
  4708 apply(simp)
       
  4709 apply(drule_tac pi="[(xa,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4710 apply(simp add: CAND_eqvt_name calc_atm)
       
  4711 apply(auto intro: CANDs_alpha)[1]
       
  4712 apply(simp)
       
  4713 apply(case_tac "za=xa")
       
  4714 apply(simp)
       
  4715 apply(drule_tac pi="[(x,xa)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4716 apply(simp add: CAND_eqvt_name calc_atm)
       
  4717 apply(auto intro: CANDs_alpha)[1]
       
  4718 apply(simp)
       
  4719 apply(drule_tac pi="[(x,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4720 apply(simp add: CAND_eqvt_name calc_atm)
       
  4721 apply(auto intro: CANDs_alpha)[1]
       
  4722 apply(drule_tac pi="[(x,za)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4723 apply(simp add: CAND_eqvt_name calc_atm)
       
  4724 apply(auto intro: CANDs_alpha)[1]
       
  4725 apply(case_tac "x=xa")
       
  4726 apply(simp)
       
  4727 apply(drule_tac pi="[(xa,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4728 apply(simp add: CAND_eqvt_name calc_atm)
       
  4729 apply(auto intro: CANDs_alpha)[1]
       
  4730 apply(simp)
       
  4731 apply(case_tac "za=xa")
       
  4732 apply(simp)
       
  4733 apply(drule_tac pi="[(x,xa)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4734 apply(simp add: CAND_eqvt_name calc_atm)
       
  4735 apply(auto intro: CANDs_alpha)[1]
       
  4736 apply(simp)
       
  4737 apply(drule_tac pi="[(x,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4738 apply(simp add: CAND_eqvt_name calc_atm)
       
  4739 apply(auto intro: CANDs_alpha)[1]
       
  4740 apply(case_tac "x=ya")
       
  4741 apply(simp)
       
  4742 apply(drule_tac pi="[(ya,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4743 apply(simp add: CAND_eqvt_name calc_atm)
       
  4744 apply(auto intro: CANDs_alpha)[1]
       
  4745 apply(simp)
       
  4746 apply(case_tac "za=ya")
       
  4747 apply(simp)
       
  4748 apply(drule_tac pi="[(x,ya)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4749 apply(simp add: CAND_eqvt_name calc_atm)
       
  4750 apply(auto intro: CANDs_alpha)[1]
       
  4751 apply(simp)
       
  4752 apply(drule_tac pi="[(x,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4753 apply(simp add: CAND_eqvt_name calc_atm)
       
  4754 apply(auto intro: CANDs_alpha)[1]
       
  4755 done
       
  4756 
       
  4757 lemmas CAND_OrL_elim = CAND_OrL_elim_aux[OF NEG_elim(2)]
       
  4758 
       
  4759 lemma CAND_AndL1_elim_aux:
       
  4760   assumes a: "(x):(AndL1 (y).M x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(AndL1 (y).M x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
       
  4761   shows "\<exists>B1 B2. B = B1 AND B2 \<and> (y):M \<in> (\<parallel>(B1)\<parallel>)" 
       
  4762 using a
       
  4763 apply(nominal_induct B rule: ty.strong_induct)
       
  4764 apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
       
  4765 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4766 apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4767 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4768 apply(case_tac "x=xa")
       
  4769 apply(simp)
       
  4770 apply(drule_tac pi="[(xa,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4771 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4772 apply(case_tac "ya=xa")
       
  4773 apply(simp)
       
  4774 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4775 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4776 apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4777 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4778 done
       
  4779 
       
  4780 lemmas CAND_AndL1_elim = CAND_AndL1_elim_aux[OF NEG_elim(2)]
       
  4781 
       
  4782 lemma CAND_AndL2_elim_aux:
       
  4783   assumes a: "(x):(AndL2 (y).M x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(AndL2 (y).M x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
       
  4784   shows "\<exists>B1 B2. B = B1 AND B2 \<and> (y):M \<in> (\<parallel>(B2)\<parallel>)" 
       
  4785 using a
       
  4786 apply(nominal_induct B rule: ty.strong_induct)
       
  4787 apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
       
  4788 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4789 apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4790 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4791 apply(case_tac "x=xa")
       
  4792 apply(simp)
       
  4793 apply(drule_tac pi="[(xa,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4794 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4795 apply(case_tac "ya=xa")
       
  4796 apply(simp)
       
  4797 apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4798 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4799 apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4800 apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
       
  4801 done
       
  4802 
       
  4803 lemmas CAND_AndL2_elim = CAND_AndL2_elim_aux[OF NEG_elim(2)]
       
  4804 
       
  4805 lemma CAND_ImpL_elim_aux:
       
  4806   assumes a: "(x):(ImpL <a>.M (z).N x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(ImpL <a>.M (z).N x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
       
  4807   shows "\<exists>B1 B2. B = B1 IMP B2 \<and> <a>:M \<in> (\<parallel><B1>\<parallel>) \<and> (z):N \<in> (\<parallel>(B2)\<parallel>)" 
       
  4808 using a
       
  4809 apply(nominal_induct B rule: ty.strong_induct)
       
  4810 apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
       
  4811 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
       
  4812 apply(drule_tac pi="[(x,y)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4813 apply(simp add: CAND_eqvt_name calc_atm)
       
  4814 apply(auto intro: CANDs_alpha)[1]
       
  4815 apply(drule_tac pi="[(x,y)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4816 apply(simp add: CAND_eqvt_name calc_atm)
       
  4817 apply(auto intro: CANDs_alpha)[1]
       
  4818 apply(drule_tac pi="[(x,y)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4819 apply(simp add: CAND_eqvt_name calc_atm)
       
  4820 apply(auto intro: CANDs_alpha)[1]
       
  4821 apply(case_tac "x=xa")
       
  4822 apply(simp)
       
  4823 apply(drule_tac pi="[(xa,y)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4824 apply(simp add: CAND_eqvt_name calc_atm)
       
  4825 apply(auto intro: CANDs_alpha)[1]
       
  4826 apply(simp)
       
  4827 apply(case_tac "y=xa")
       
  4828 apply(simp)
       
  4829 apply(drule_tac pi="[(x,xa)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4830 apply(simp add: CAND_eqvt_name calc_atm)
       
  4831 apply(auto intro: CANDs_alpha)[1]
       
  4832 apply(simp)
       
  4833 apply(drule_tac pi="[(x,y)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4834 apply(simp add: CAND_eqvt_name calc_atm)
       
  4835 apply(auto intro: CANDs_alpha)[1]
       
  4836 done
       
  4837 
       
  4838 lemmas CAND_ImpL_elim = CAND_ImpL_elim_aux[OF NEG_elim(2)]
       
  4839 
       
  4840 lemma CAND_ImpR_elim:
       
  4841   assumes a: "<a>:ImpR (x).<b>.M a \<in> (\<parallel><B>\<parallel>)" "<a>:ImpR (x).<b>.M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  4842   shows "\<exists>B1 B2. B = B1 IMP B2 \<and> 
       
  4843                  (\<forall>z P. x\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B2)\<parallel> \<longrightarrow> (x):(M{b:=(z).P}) \<in> \<parallel>(B1)\<parallel>) \<and>
       
  4844                  (\<forall>c Q. b\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><B1>\<parallel> \<longrightarrow> <b>:(M{x:=<c>.Q}) \<in> \<parallel><B2>\<parallel>)" 
       
  4845 using a
       
  4846 apply(nominal_induct B rule: ty.strong_induct)
       
  4847 apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
       
  4848 apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm fresh_prod fresh_bij)
       
  4849 apply(generate_fresh "name") 
       
  4850 apply(generate_fresh "coname")
       
  4851 apply(drule_tac a="ca" and z="c" in alpha_name_coname)
       
  4852 apply(simp) 
       
  4853 apply(simp) 
       
  4854 apply(simp) 
       
  4855 apply(drule_tac x="[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
       
  4856 apply(drule_tac x="[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
       
  4857 apply(drule mp)
       
  4858 apply(rule conjI)
       
  4859 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  4860 apply(rule conjI)
       
  4861 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  4862 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4863 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4864 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4865 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4866 apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4867 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4868 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4869 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4870 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4871 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4872 apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4873 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4874 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4875 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4876 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4877 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4878 apply(generate_fresh "name")
       
  4879 apply(generate_fresh "coname")
       
  4880 apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
       
  4881 apply(simp)
       
  4882 apply(simp)
       
  4883 apply(simp)
       
  4884 apply(drule_tac x="[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
       
  4885 apply(drule_tac x="[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
       
  4886 apply(drule mp)
       
  4887 apply(rule conjI)
       
  4888 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  4889 apply(rule conjI)
       
  4890 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  4891 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4892 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4893 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4894 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4895 apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4896 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4897 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4898 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4899 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4900 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4901 apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4902 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4903 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4904 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4905 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4906 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4907 apply(generate_fresh "name")
       
  4908 apply(generate_fresh "coname")
       
  4909 apply(drule_tac a="ca" and z="c" in alpha_name_coname)
       
  4910 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  4911 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  4912 apply(auto)[1]
       
  4913 apply(simp)
       
  4914 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
       
  4915 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
       
  4916 apply(drule mp)
       
  4917 apply(rule conjI)
       
  4918 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  4919 apply(rule conjI)
       
  4920 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  4921 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4922 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4923 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4924 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4925 apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4926 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4927 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4928 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4929 apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4930 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4931 apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4932 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4933 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4934 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4935 apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4936 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4937 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4938 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4939 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4940 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4941 apply(generate_fresh "name")
       
  4942 apply(generate_fresh "coname")
       
  4943 apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
       
  4944 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  4945 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  4946 apply(auto)[1]
       
  4947 apply(simp)
       
  4948 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
       
  4949 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
       
  4950 apply(drule mp)
       
  4951 apply(rule conjI)
       
  4952 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  4953 apply(rule conjI)
       
  4954 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  4955 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4956 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4957 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4958 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4959 apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4960 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4961 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4962 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4963 apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4964 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4965 apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4966 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4967 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4968 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4969 apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4970 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4971 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4972 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4973 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4974 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  4975 apply(case_tac "a=aa")
       
  4976 apply(simp)
       
  4977 apply(generate_fresh "name")
       
  4978 apply(generate_fresh "coname")
       
  4979 apply(drule_tac a="ca" and z="c" in alpha_name_coname)
       
  4980 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  4981 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  4982 apply(auto)[1]
       
  4983 apply(simp)
       
  4984 apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
       
  4985 apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
       
  4986 apply(drule mp)
       
  4987 apply(rule conjI)
       
  4988 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  4989 apply(rule conjI)
       
  4990 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  4991 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4992 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4993 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4994 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4995 apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  4996 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4997 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  4998 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  4999 apply(drule_tac pi="[(aa,ba)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5000 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5001 apply(drule_tac pi="[(aa,ba)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5002 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5003 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5004 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5005 apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5006 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5007 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5008 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5009 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5010 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5011 apply(simp)
       
  5012 apply(case_tac "ba=aa")
       
  5013 apply(simp)
       
  5014 apply(generate_fresh "name")
       
  5015 apply(generate_fresh "coname")
       
  5016 apply(drule_tac a="ca" and z="c" in alpha_name_coname)
       
  5017 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5018 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5019 apply(auto)[1]
       
  5020 apply(simp)
       
  5021 apply(drule_tac x="[(a,aa)]\<bullet>[(xa,c)]\<bullet>[(a,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
       
  5022 apply(drule_tac x="[(a,aa)]\<bullet>[(xa,c)]\<bullet>[(a,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
       
  5023 apply(drule mp)
       
  5024 apply(rule conjI)
       
  5025 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  5026 apply(rule conjI)
       
  5027 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  5028 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5029 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5030 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5031 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5032 apply(drule_tac pi="[(a,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5033 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5034 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5035 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5036 apply(drule_tac pi="[(a,aa)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5037 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5038 apply(drule_tac pi="[(a,aa)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5039 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5040 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5041 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5042 apply(drule_tac pi="[(a,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5043 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5044 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5045 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5046 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5047 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5048 apply(simp)
       
  5049 apply(generate_fresh "name")
       
  5050 apply(generate_fresh "coname")
       
  5051 apply(drule_tac a="ca" and z="c" in alpha_name_coname)
       
  5052 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5053 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5054 apply(auto)[1]
       
  5055 apply(simp)
       
  5056 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
       
  5057 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
       
  5058 apply(drule mp)
       
  5059 apply(rule conjI)
       
  5060 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  5061 apply(rule conjI)
       
  5062 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  5063 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5064 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5065 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5066 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5067 apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5068 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5069 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5070 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5071 apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5072 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5073 apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5074 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5075 apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5076 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5077 apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5078 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5079 apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5080 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5081 apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5082 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5083 apply(case_tac "a=aa")
       
  5084 apply(simp)
       
  5085 apply(generate_fresh "name")
       
  5086 apply(generate_fresh "coname")
       
  5087 apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
       
  5088 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5089 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5090 apply(auto)[1]
       
  5091 apply(simp)
       
  5092 apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
       
  5093 apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
       
  5094 apply(drule mp)
       
  5095 apply(rule conjI)
       
  5096 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  5097 apply(rule conjI)
       
  5098 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  5099 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5100 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5101 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5102 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5103 apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5104 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5105 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5106 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5107 apply(drule_tac pi="[(aa,ba)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5108 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5109 apply(drule_tac pi="[(aa,ba)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5110 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5111 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5112 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5113 apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5114 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5115 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5116 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5117 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5118 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5119 apply(simp)
       
  5120 apply(case_tac "ba=aa")
       
  5121 apply(simp)
       
  5122 apply(generate_fresh "name")
       
  5123 apply(generate_fresh "coname")
       
  5124 apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
       
  5125 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5126 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5127 apply(auto)[1]
       
  5128 apply(simp)
       
  5129 apply(drule_tac x="[(a,aa)]\<bullet>[(xa,ca)]\<bullet>[(a,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
       
  5130 apply(drule_tac x="[(a,aa)]\<bullet>[(xa,ca)]\<bullet>[(a,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
       
  5131 apply(drule mp)
       
  5132 apply(rule conjI)
       
  5133 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  5134 apply(rule conjI)
       
  5135 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  5136 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5137 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5138 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5139 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5140 apply(drule_tac pi="[(a,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5141 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5142 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5143 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5144 apply(drule_tac pi="[(a,aa)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5145 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5146 apply(drule_tac pi="[(a,aa)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5147 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5148 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5149 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5150 apply(drule_tac pi="[(a,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5151 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5152 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5153 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5154 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5155 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5156 apply(simp)
       
  5157 apply(generate_fresh "name")
       
  5158 apply(generate_fresh "coname")
       
  5159 apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
       
  5160 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5161 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
       
  5162 apply(auto)[1]
       
  5163 apply(simp)
       
  5164 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
       
  5165 apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
       
  5166 apply(drule mp)
       
  5167 apply(rule conjI)
       
  5168 apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
       
  5169 apply(rule conjI)
       
  5170 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
       
  5171 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5172 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5173 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5174 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5175 apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5176 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5177 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5178 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5179 apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5180 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5181 apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5182 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5183 apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5184 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5185 apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5186 apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5187 apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
       
  5188 apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
       
  5189 apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
       
  5190 apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
       
  5191 done
       
  5192 
       
  5193 text {* Main lemma 1 *}
       
  5194 
       
  5195 lemma AXIOMS_imply_SNa:
       
  5196   shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> SNa M"
       
  5197   and   "(x):M \<in> AXIOMSn B \<Longrightarrow> SNa M"
       
  5198 apply -
       
  5199 apply(auto simp add: AXIOMSn_def AXIOMSc_def ntrm.inject ctrm.inject alpha)
       
  5200 apply(rule Ax_in_SNa)+
       
  5201 done
       
  5202 
       
  5203 lemma BINDING_imply_SNa:
       
  5204   shows "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa M"
       
  5205   and   "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> SNa M"
       
  5206 apply -
       
  5207 apply(auto simp add: BINDINGn_def BINDINGc_def ntrm.inject ctrm.inject alpha)
       
  5208 apply(drule_tac x="x" in spec)
       
  5209 apply(drule_tac x="Ax x a" in spec)
       
  5210 apply(drule mp)
       
  5211 apply(rule Ax_in_CANDs)
       
  5212 apply(drule a_star_preserves_SNa)
       
  5213 apply(rule subst_with_ax2)
       
  5214 apply(simp add: crename_id)
       
  5215 apply(drule_tac x="x" in spec)
       
  5216 apply(drule_tac x="Ax x aa" in spec)
       
  5217 apply(drule mp)
       
  5218 apply(rule Ax_in_CANDs)
       
  5219 apply(drule a_star_preserves_SNa)
       
  5220 apply(rule subst_with_ax2)
       
  5221 apply(simp add: crename_id SNa_eqvt)
       
  5222 apply(drule_tac x="a" in spec)
       
  5223 apply(drule_tac x="Ax x a" in spec)
       
  5224 apply(drule mp)
       
  5225 apply(rule Ax_in_CANDs)
       
  5226 apply(drule a_star_preserves_SNa)
       
  5227 apply(rule subst_with_ax1)
       
  5228 apply(simp add: nrename_id)
       
  5229 apply(drule_tac x="a" in spec)
       
  5230 apply(drule_tac x="Ax xa a" in spec)
       
  5231 apply(drule mp)
       
  5232 apply(rule Ax_in_CANDs)
       
  5233 apply(drule a_star_preserves_SNa)
       
  5234 apply(rule subst_with_ax1)
       
  5235 apply(simp add: nrename_id SNa_eqvt)
       
  5236 done
       
  5237 
       
  5238 lemma CANDs_imply_SNa:
       
  5239   shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M"
       
  5240   and   "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M"
       
  5241 proof(induct B arbitrary: a x M rule: ty.induct)
       
  5242   case (PR X)
       
  5243   { case 1 
       
  5244     have "<a>:M \<in> \<parallel><PR X>\<parallel>" by fact
       
  5245     then have "<a>:M \<in> NEGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
       
  5246     then have "<a>:M \<in> AXIOMSc (PR X) \<union> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
       
  5247     moreover
       
  5248     { assume "<a>:M \<in> AXIOMSc (PR X)"
       
  5249       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5250     }
       
  5251     moreover
       
  5252     { assume "<a>:M \<in> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)"
       
  5253       then have "SNa M" by (simp add: BINDING_imply_SNa)
       
  5254     }
       
  5255     ultimately show "SNa M" by blast 
       
  5256   next
       
  5257     case 2
       
  5258     have "(x):M \<in> (\<parallel>(PR X)\<parallel>)" by fact
       
  5259     then have "(x):M \<in> NEGn (PR X) (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
       
  5260     then have "(x):M \<in> AXIOMSn (PR X) \<union> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" by simp
       
  5261     moreover
       
  5262     { assume "(x):M \<in> AXIOMSn (PR X)"
       
  5263       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5264     }
       
  5265     moreover
       
  5266     { assume "(x):M \<in> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)"
       
  5267       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5268     }
       
  5269     ultimately show "SNa M" by blast
       
  5270   }
       
  5271 next
       
  5272   case (NOT B)
       
  5273   have ih1: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
       
  5274   have ih2: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
       
  5275   { case 1
       
  5276     have "<a>:M \<in> (\<parallel><NOT B>\<parallel>)" by fact
       
  5277     then have "<a>:M \<in> NEGc (NOT B) (\<parallel>(NOT B)\<parallel>)" by simp
       
  5278     then have "<a>:M \<in> AXIOMSc (NOT B) \<union> BINDINGc (NOT B) (\<parallel>(NOT B)\<parallel>) \<union> NOTRIGHT (NOT B) (\<parallel>(B)\<parallel>)" by simp
       
  5279      moreover
       
  5280     { assume "<a>:M \<in> AXIOMSc (NOT B)"
       
  5281       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5282     }
       
  5283     moreover
       
  5284     { assume "<a>:M \<in> BINDINGc (NOT B) (\<parallel>(NOT B)\<parallel>)"
       
  5285       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5286     }
       
  5287      moreover
       
  5288     { assume "<a>:M \<in> NOTRIGHT (NOT B) (\<parallel>(B)\<parallel>)"
       
  5289       then obtain x' M' where eq: "M = NotR (x').M' a" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
       
  5290         using NOTRIGHT_elim by blast
       
  5291       then have "SNa M'" using ih2 by blast
       
  5292       then have "SNa M" using eq by (simp add: NotR_in_SNa)
       
  5293     }
       
  5294     ultimately show "SNa M" by blast
       
  5295   next
       
  5296     case 2
       
  5297     have "(x):M \<in> (\<parallel>(NOT B)\<parallel>)" by fact
       
  5298     then have "(x):M \<in> NEGn (NOT B) (\<parallel><NOT B>\<parallel>)" using NEG_simp by blast
       
  5299     then have "(x):M \<in> AXIOMSn (NOT B) \<union> BINDINGn (NOT B) (\<parallel><NOT B>\<parallel>) \<union> NOTLEFT (NOT B) (\<parallel><B>\<parallel>)" 
       
  5300       by (simp only: NEGn.simps)
       
  5301      moreover
       
  5302     { assume "(x):M \<in> AXIOMSn (NOT B)"
       
  5303       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5304     }
       
  5305     moreover
       
  5306     { assume "(x):M \<in> BINDINGn (NOT B) (\<parallel><NOT B>\<parallel>)"
       
  5307       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5308     }
       
  5309      moreover
       
  5310     { assume "(x):M \<in> NOTLEFT (NOT B) (\<parallel><B>\<parallel>)"
       
  5311       then obtain a' M' where eq: "M = NotL <a'>.M' x" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)"
       
  5312         using NOTLEFT_elim by blast
       
  5313       then have "SNa M'" using ih1 by blast
       
  5314       then have "SNa M" using eq by (simp add: NotL_in_SNa)
       
  5315     }
       
  5316     ultimately show "SNa M" by blast
       
  5317   }
       
  5318 next
       
  5319   case (AND A B)
       
  5320   have ih1: "\<And>a M. <a>:M \<in> \<parallel><A>\<parallel> \<Longrightarrow> SNa M" by fact
       
  5321   have ih2: "\<And>x M. (x):M \<in> \<parallel>(A)\<parallel> \<Longrightarrow> SNa M" by fact
       
  5322   have ih3: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
       
  5323   have ih4: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
       
  5324   { case 1
       
  5325     have "<a>:M \<in> (\<parallel><A AND B>\<parallel>)" by fact
       
  5326     then have "<a>:M \<in> NEGc (A AND B) (\<parallel>(A AND B)\<parallel>)" by simp
       
  5327     then have "<a>:M \<in> AXIOMSc (A AND B) \<union> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>) 
       
  5328                                   \<union> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" by simp
       
  5329      moreover
       
  5330     { assume "<a>:M \<in> AXIOMSc (A AND B)"
       
  5331       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5332     }
       
  5333     moreover
       
  5334     { assume "<a>:M \<in> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)"
       
  5335       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5336     }
       
  5337      moreover
       
  5338     { assume "<a>:M \<in> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)"
       
  5339       then obtain a' M' b' N' where eq: "M = AndR <a'>.M' <b'>.N' a" 
       
  5340                                 and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" and "<b'>:N' \<in> (\<parallel><B>\<parallel>)"
       
  5341         by (erule_tac ANDRIGHT_elim, blast)
       
  5342       then have "SNa M'" and "SNa N'" using ih1 ih3 by blast+
       
  5343       then have "SNa M" using eq by (simp add: AndR_in_SNa)
       
  5344     }
       
  5345     ultimately show "SNa M" by blast
       
  5346   next
       
  5347     case 2
       
  5348     have "(x):M \<in> (\<parallel>(A AND B)\<parallel>)" by fact
       
  5349     then have "(x):M \<in> NEGn (A AND B) (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
       
  5350     then have "(x):M \<in> AXIOMSn (A AND B) \<union> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>) 
       
  5351                        \<union> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>) \<union> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" 
       
  5352       by (simp only: NEGn.simps)
       
  5353      moreover
       
  5354     { assume "(x):M \<in> AXIOMSn (A AND B)"
       
  5355       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5356     }
       
  5357     moreover
       
  5358     { assume "(x):M \<in> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)"
       
  5359       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5360     }
       
  5361      moreover
       
  5362     { assume "(x):M \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)"
       
  5363       then obtain x' M' where eq: "M = AndL1 (x').M' x" and "(x'):M' \<in> (\<parallel>(A)\<parallel>)"
       
  5364         using ANDLEFT1_elim by blast
       
  5365       then have "SNa M'" using ih2 by blast
       
  5366       then have "SNa M" using eq by (simp add: AndL1_in_SNa)
       
  5367     }
       
  5368     moreover
       
  5369     { assume "(x):M \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)"
       
  5370       then obtain x' M' where eq: "M = AndL2 (x').M' x" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
       
  5371         using ANDLEFT2_elim by blast
       
  5372       then have "SNa M'" using ih4 by blast
       
  5373       then have "SNa M" using eq by (simp add: AndL2_in_SNa)
       
  5374     }
       
  5375     ultimately show "SNa M" by blast
       
  5376   }
       
  5377 next
       
  5378   case (OR A B)
       
  5379   have ih1: "\<And>a M. <a>:M \<in> \<parallel><A>\<parallel> \<Longrightarrow> SNa M" by fact
       
  5380   have ih2: "\<And>x M. (x):M \<in> \<parallel>(A)\<parallel> \<Longrightarrow> SNa M" by fact
       
  5381   have ih3: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
       
  5382   have ih4: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
       
  5383   { case 1
       
  5384     have "<a>:M \<in> (\<parallel><A OR B>\<parallel>)" by fact
       
  5385     then have "<a>:M \<in> NEGc (A OR B) (\<parallel>(A OR B)\<parallel>)" by simp
       
  5386     then have "<a>:M \<in> AXIOMSc (A OR B) \<union> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>) 
       
  5387                                   \<union> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>) \<union> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" by simp
       
  5388      moreover
       
  5389     { assume "<a>:M \<in> AXIOMSc (A OR B)"
       
  5390       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5391     }
       
  5392     moreover
       
  5393     { assume "<a>:M \<in> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)"
       
  5394       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5395     }
       
  5396      moreover
       
  5397     { assume "<a>:M \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)"
       
  5398       then obtain a' M' where eq: "M = OrR1 <a'>.M' a" 
       
  5399                                 and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" 
       
  5400         by (erule_tac ORRIGHT1_elim, blast)
       
  5401       then have "SNa M'" using ih1 by blast
       
  5402       then have "SNa M" using eq by (simp add: OrR1_in_SNa)
       
  5403     }
       
  5404      moreover
       
  5405     { assume "<a>:M \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)"
       
  5406       then obtain a' M' where eq: "M = OrR2 <a'>.M' a" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)" 
       
  5407         using ORRIGHT2_elim by blast
       
  5408       then have "SNa M'" using ih3 by blast
       
  5409       then have "SNa M" using eq by (simp add: OrR2_in_SNa)
       
  5410     }
       
  5411     ultimately show "SNa M" by blast
       
  5412   next
       
  5413     case 2
       
  5414     have "(x):M \<in> (\<parallel>(A OR B)\<parallel>)" by fact
       
  5415     then have "(x):M \<in> NEGn (A OR B) (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
       
  5416     then have "(x):M \<in> AXIOMSn (A OR B) \<union> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>) 
       
  5417                        \<union> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" 
       
  5418       by (simp only: NEGn.simps)
       
  5419      moreover
       
  5420     { assume "(x):M \<in> AXIOMSn (A OR B)"
       
  5421       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5422     }
       
  5423     moreover
       
  5424     { assume "(x):M \<in> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)"
       
  5425       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5426     }
       
  5427      moreover
       
  5428     { assume "(x):M \<in> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)"
       
  5429       then obtain x' M' y' N' where eq: "M = OrL (x').M' (y').N' x" 
       
  5430                                 and "(x'):M' \<in> (\<parallel>(A)\<parallel>)" and  "(y'):N' \<in> (\<parallel>(B)\<parallel>)"
       
  5431         by (erule_tac ORLEFT_elim, blast)
       
  5432       then have "SNa M'" and "SNa N'" using ih2 ih4 by blast+
       
  5433       then have "SNa M" using eq by (simp add: OrL_in_SNa)
       
  5434     }
       
  5435     ultimately show "SNa M" by blast
       
  5436   }
       
  5437 next 
       
  5438   case (IMP A B)
       
  5439   have ih1: "\<And>a M. <a>:M \<in> \<parallel><A>\<parallel> \<Longrightarrow> SNa M" by fact
       
  5440   have ih2: "\<And>x M. (x):M \<in> \<parallel>(A)\<parallel> \<Longrightarrow> SNa M" by fact
       
  5441   have ih3: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
       
  5442   have ih4: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
       
  5443   { case 1
       
  5444     have "<a>:M \<in> (\<parallel><A IMP B>\<parallel>)" by fact
       
  5445     then have "<a>:M \<in> NEGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" by simp
       
  5446     then have "<a>:M \<in> AXIOMSc (A IMP B) \<union> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>) 
       
  5447                                   \<union> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" by simp
       
  5448      moreover
       
  5449     { assume "<a>:M \<in> AXIOMSc (A IMP B)"
       
  5450       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5451     }
       
  5452     moreover
       
  5453     { assume "<a>:M \<in> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)"
       
  5454       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5455     }
       
  5456      moreover
       
  5457     { assume "<a>:M \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)"
       
  5458       then obtain x' a' M' where eq: "M = ImpR (x').<a'>.M' a" 
       
  5459                            and imp: "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(M'{a':=(z).P}) \<in> \<parallel>(A)\<parallel>"    
       
  5460         by (erule_tac IMPRIGHT_elim, blast)
       
  5461       obtain z::"name" where fs: "z\<sharp>x'" by (rule_tac exists_fresh, rule fin_supp, blast)
       
  5462       have "(z):Ax z a'\<in> \<parallel>(B)\<parallel>" by (simp add: Ax_in_CANDs)
       
  5463       with imp fs have "(x'):(M'{a':=(z).Ax z a'}) \<in> \<parallel>(A)\<parallel>" by (simp add: fresh_prod fresh_atm)
       
  5464       then have "SNa (M'{a':=(z).Ax z a'})" using ih2 by blast
       
  5465       moreover 
       
  5466       have "M'{a':=(z).Ax z a'} \<longrightarrow>\<^isub>a* M'[a'\<turnstile>c>a']" by (simp add: subst_with_ax2)
       
  5467       ultimately have "SNa (M'[a'\<turnstile>c>a'])" by (simp add: a_star_preserves_SNa)
       
  5468       then have "SNa M'" by (simp add: crename_id)
       
  5469       then have "SNa M" using eq by (simp add: ImpR_in_SNa)
       
  5470     }
       
  5471     ultimately show "SNa M" by blast
       
  5472   next
       
  5473     case 2
       
  5474     have "(x):M \<in> (\<parallel>(A IMP B)\<parallel>)" by fact
       
  5475     then have "(x):M \<in> NEGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
       
  5476     then have "(x):M \<in> AXIOMSn (A IMP B) \<union> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>) 
       
  5477                        \<union> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" 
       
  5478       by (simp only: NEGn.simps)
       
  5479      moreover
       
  5480     { assume "(x):M \<in> AXIOMSn (A IMP B)"
       
  5481       then have "SNa M" by (simp add: AXIOMS_imply_SNa)
       
  5482     }
       
  5483     moreover
       
  5484     { assume "(x):M \<in> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)"
       
  5485       then have "SNa M" by (simp only: BINDING_imply_SNa)
       
  5486     }
       
  5487      moreover
       
  5488     { assume "(x):M \<in> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)"
       
  5489       then obtain a' M' y' N' where eq: "M = ImpL <a'>.M' (y').N' x" 
       
  5490                                 and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" and  "(y'):N' \<in> (\<parallel>(B)\<parallel>)"
       
  5491         by (erule_tac IMPLEFT_elim, blast)
       
  5492       then have "SNa M'" and "SNa N'" using ih1 ih4 by blast+
       
  5493       then have "SNa M" using eq by (simp add: ImpL_in_SNa)
       
  5494     }
       
  5495     ultimately show "SNa M" by blast
       
  5496   }
       
  5497 qed 
       
  5498 
       
  5499 text {* Main lemma 2 *}
       
  5500 
       
  5501 lemma AXIOMS_preserved:
       
  5502   shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> AXIOMSc B"
       
  5503   and   "(x):M \<in> AXIOMSn B \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> AXIOMSn B"
       
  5504   apply(simp_all add: AXIOMSc_def AXIOMSn_def)
       
  5505   apply(auto simp add: ntrm.inject ctrm.inject alpha)
       
  5506   apply(drule ax_do_not_a_star_reduce)
       
  5507   apply(auto)
       
  5508   apply(drule ax_do_not_a_star_reduce)
       
  5509   apply(auto)
       
  5510   apply(drule ax_do_not_a_star_reduce)
       
  5511   apply(auto)
       
  5512   apply(drule ax_do_not_a_star_reduce)
       
  5513   apply(auto)
       
  5514   done  
       
  5515 
       
  5516 lemma BINDING_preserved:
       
  5517   shows "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  5518   and   "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> BINDINGn B (\<parallel><B>\<parallel>)"
       
  5519 proof -
       
  5520   assume red: "M \<longrightarrow>\<^isub>a* M'"
       
  5521   assume asm: "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  5522   {
       
  5523     fix x::"name" and  P::"trm"
       
  5524     from asm have "((x):P) \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa (M{a:=(x).P})" by (simp add: BINDINGc_elim)
       
  5525     moreover
       
  5526     have "M{a:=(x).P} \<longrightarrow>\<^isub>a* M'{a:=(x).P}" using red by (simp add: a_star_subst2)
       
  5527     ultimately 
       
  5528     have "((x):P) \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa (M'{a:=(x).P})" by (simp add: a_star_preserves_SNa)
       
  5529   }
       
  5530   then show "<a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)" by (auto simp add: BINDINGc_def)
       
  5531 next
       
  5532   assume red: "M \<longrightarrow>\<^isub>a* M'"
       
  5533   assume asm: "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
       
  5534   {
       
  5535     fix c::"coname" and  P::"trm"
       
  5536     from asm have "(<c>:P) \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> SNa (M{x:=<c>.P})" by (simp add: BINDINGn_elim)
       
  5537     moreover
       
  5538     have "M{x:=<c>.P} \<longrightarrow>\<^isub>a* M'{x:=<c>.P}" using red by (simp add: a_star_subst1)
       
  5539     ultimately 
       
  5540     have "(<c>:P) \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> SNa (M'{x:=<c>.P})" by (simp add: a_star_preserves_SNa)
       
  5541   }
       
  5542   then show "(x):M' \<in> BINDINGn B (\<parallel><B>\<parallel>)" by (auto simp add: BINDINGn_def)
       
  5543 qed
       
  5544     
       
  5545 lemma CANDs_preserved:
       
  5546   shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
       
  5547   and   "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" 
       
  5548 proof(nominal_induct B arbitrary: a x M M' rule: ty.strong_induct) 
       
  5549   case (PR X)
       
  5550   { case 1 
       
  5551     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5552     have "<a>:M \<in> \<parallel><PR X>\<parallel>" by fact
       
  5553     then have "<a>:M \<in> NEGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
       
  5554     then have "<a>:M \<in> AXIOMSc (PR X) \<union> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
       
  5555     moreover
       
  5556     { assume "<a>:M \<in> AXIOMSc (PR X)"
       
  5557       then have "<a>:M' \<in> AXIOMSc (PR X)" using asm by (simp only: AXIOMS_preserved)
       
  5558     }
       
  5559     moreover
       
  5560     { assume "<a>:M \<in> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)"
       
  5561       then have "<a>:M' \<in> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" using asm by (simp add: BINDING_preserved)
       
  5562     }
       
  5563     ultimately have "<a>:M' \<in> AXIOMSc (PR X) \<union> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" by blast
       
  5564     then have "<a>:M' \<in> NEGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
       
  5565     then show "<a>:M' \<in> (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
       
  5566   next
       
  5567     case 2
       
  5568     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5569     have "(x):M \<in> \<parallel>(PR X)\<parallel>" by fact
       
  5570     then have "(x):M \<in> NEGn (PR X) (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
       
  5571     then have "(x):M \<in> AXIOMSn (PR X) \<union> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" by simp
       
  5572     moreover
       
  5573     { assume "(x):M \<in> AXIOMSn (PR X)"
       
  5574       then have "(x):M' \<in> AXIOMSn (PR X)" using asm by (simp only: AXIOMS_preserved) 
       
  5575     }
       
  5576     moreover
       
  5577     { assume "(x):M \<in> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)"
       
  5578       then have "(x):M' \<in> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5579     }
       
  5580     ultimately have "(x):M' \<in> AXIOMSn (PR X) \<union> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" by blast
       
  5581     then have "(x):M' \<in> NEGn (PR X) (\<parallel><PR X>\<parallel>)" by simp
       
  5582     then show "(x):M' \<in> (\<parallel>(PR X)\<parallel>)" using NEG_simp by blast
       
  5583   }
       
  5584 next
       
  5585   case (IMP A B)
       
  5586   have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
       
  5587   have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
       
  5588   have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
       
  5589   have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
       
  5590   { case 1 
       
  5591     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5592     have "<a>:M \<in> \<parallel><A IMP B>\<parallel>" by fact
       
  5593     then have "<a>:M \<in> NEGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" by simp
       
  5594     then have "<a>:M \<in> AXIOMSc (A IMP B) \<union> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>) 
       
  5595                             \<union> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" by simp
       
  5596     moreover
       
  5597     { assume "<a>:M \<in> AXIOMSc (A IMP B)"
       
  5598       then have "<a>:M' \<in> AXIOMSc (A IMP B)" using asm by (simp only: AXIOMS_preserved)
       
  5599     }
       
  5600     moreover
       
  5601     { assume "<a>:M \<in> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)"
       
  5602       then have "<a>:M' \<in> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5603     }
       
  5604     moreover
       
  5605     { assume "<a>:M \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)"
       
  5606       then obtain x' a' N' where eq: "M = ImpR (x').<a'>.N' a" and fic: "fic (ImpR (x').<a'>.N' a) a"
       
  5607                            and imp1: "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N'{a':=(z).P}) \<in> \<parallel>(A)\<parallel>" 
       
  5608                            and imp2: "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><A>\<parallel> \<longrightarrow> <a'>:(N'{x':=<c>.Q}) \<in> \<parallel><B>\<parallel>"
       
  5609         using IMPRIGHT_elim by blast
       
  5610       from eq asm obtain N'' where eq': "M' = ImpR (x').<a'>.N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''" 
       
  5611         using a_star_redu_ImpR_elim by (blast)
       
  5612       from imp1 have "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N''{a':=(z).P}) \<in> \<parallel>(A)\<parallel>" using red ih2
       
  5613         apply(auto)
       
  5614         apply(drule_tac x="z" in spec)
       
  5615         apply(drule_tac x="P" in spec)
       
  5616         apply(simp)
       
  5617         apply(drule_tac a_star_subst2)
       
  5618         apply(blast)
       
  5619         done
       
  5620       moreover
       
  5621       from imp2 have "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><A>\<parallel> \<longrightarrow> <a'>:(N''{x':=<c>.Q}) \<in> \<parallel><B>\<parallel>" using red ih3
       
  5622         apply(auto)
       
  5623         apply(drule_tac x="c" in spec)
       
  5624         apply(drule_tac x="Q" in spec)
       
  5625         apply(simp)
       
  5626         apply(drule_tac a_star_subst1)
       
  5627         apply(blast)
       
  5628         done
       
  5629       moreover
       
  5630       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       
  5631       ultimately have "<a>:M' \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" using eq' by auto
       
  5632     }
       
  5633     ultimately have "<a>:M' \<in> AXIOMSc (A IMP B) \<union> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)
       
  5634                                             \<union> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" by blast
       
  5635     then have "<a>:M' \<in> NEGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" by simp
       
  5636     then show "<a>:M' \<in> (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
       
  5637   next
       
  5638     case 2
       
  5639     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5640     have "(x):M \<in> \<parallel>(A IMP B)\<parallel>" by fact
       
  5641     then have "(x):M \<in> NEGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
       
  5642     then have "(x):M \<in> AXIOMSn (A IMP B) \<union> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>) 
       
  5643                                               \<union> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" by simp
       
  5644     moreover
       
  5645     { assume "(x):M \<in> AXIOMSn (A IMP B)"
       
  5646       then have "(x):M' \<in> AXIOMSn (A IMP B)" using asm by (simp only: AXIOMS_preserved)
       
  5647     }
       
  5648     moreover
       
  5649     { assume "(x):M \<in> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)"
       
  5650       then have "(x):M' \<in> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5651     }
       
  5652     moreover
       
  5653     { assume "(x):M \<in> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)"
       
  5654       then obtain a' T' y' N' where eq: "M = ImpL <a'>.T' (y').N' x" 
       
  5655                              and fin: "fin (ImpL <a'>.T' (y').N' x) x"
       
  5656                              and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "(y'):N' \<in> \<parallel>(B)\<parallel>"
       
  5657         by (erule_tac IMPLEFT_elim, blast)
       
  5658       from eq asm obtain T'' N'' where eq': "M' = ImpL <a'>.T'' (y').N'' x" 
       
  5659                                  and red1: "T' \<longrightarrow>\<^isub>a* T''"  and red2: "N' \<longrightarrow>\<^isub>a* N''"
       
  5660         using a_star_redu_ImpL_elim by blast
       
  5661       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       
  5662       moreover
       
  5663       from imp1 red1 have "<a'>:T'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
       
  5664       moreover
       
  5665       from imp2 red2 have "(y'):N'' \<in> \<parallel>(B)\<parallel>" using ih4 by simp
       
  5666       ultimately have "(x):M' \<in> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" using eq' by (simp, blast) 
       
  5667     }
       
  5668     ultimately have "(x):M' \<in> AXIOMSn (A IMP B) \<union> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)
       
  5669                                               \<union> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" by blast
       
  5670     then have "(x):M' \<in> NEGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" by simp
       
  5671     then show "(x):M' \<in> (\<parallel>(A IMP B)\<parallel>)" using NEG_simp by blast
       
  5672   }
       
  5673 next
       
  5674   case (AND A B)
       
  5675   have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
       
  5676   have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
       
  5677   have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
       
  5678   have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
       
  5679   { case 1 
       
  5680     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5681     have "<a>:M \<in> \<parallel><A AND B>\<parallel>" by fact
       
  5682     then have "<a>:M \<in> NEGc (A AND B) (\<parallel>(A AND B)\<parallel>)" by simp
       
  5683     then have "<a>:M \<in> AXIOMSc (A AND B) \<union> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>) 
       
  5684                                               \<union> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" by simp
       
  5685     moreover
       
  5686     { assume "<a>:M \<in> AXIOMSc (A AND B)"
       
  5687       then have "<a>:M' \<in> AXIOMSc (A AND B)" using asm by (simp only: AXIOMS_preserved)
       
  5688     }
       
  5689     moreover
       
  5690     { assume "<a>:M \<in> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)"
       
  5691       then have "<a>:M' \<in> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5692     }
       
  5693     moreover
       
  5694     { assume "<a>:M \<in> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)"
       
  5695       then obtain a' T' b' N' where eq: "M = AndR <a'>.T' <b'>.N' a" 
       
  5696                               and fic: "fic (AndR <a'>.T' <b'>.N' a) a"
       
  5697                            and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "<b'>:N' \<in> \<parallel><B>\<parallel>"
       
  5698         using ANDRIGHT_elim by blast
       
  5699       from eq asm obtain T'' N'' where eq': "M' = AndR <a'>.T'' <b'>.N'' a" 
       
  5700                           and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''" 
       
  5701         using a_star_redu_AndR_elim by blast
       
  5702       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       
  5703       moreover
       
  5704       from imp1 red1 have "<a'>:T'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
       
  5705       moreover
       
  5706       from imp2 red2 have "<b'>:N'' \<in> \<parallel><B>\<parallel>" using ih3 by simp
       
  5707       ultimately have "<a>:M' \<in> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" using eq' by (simp, blast) 
       
  5708     }
       
  5709     ultimately have "<a>:M' \<in> AXIOMSc (A AND B) \<union> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)
       
  5710                                               \<union> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" by blast
       
  5711     then have "<a>:M' \<in> NEGc (A AND B) (\<parallel>(A AND B)\<parallel>)" by simp
       
  5712     then show "<a>:M' \<in> (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
       
  5713   next
       
  5714     case 2
       
  5715     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5716     have "(x):M \<in> \<parallel>(A AND B)\<parallel>" by fact
       
  5717     then have "(x):M \<in> NEGn (A AND B) (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
       
  5718     then have "(x):M \<in> AXIOMSn (A AND B) \<union> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>) 
       
  5719                                      \<union> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>) \<union> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" by simp
       
  5720     moreover
       
  5721     { assume "(x):M \<in> AXIOMSn (A AND B)"
       
  5722       then have "(x):M' \<in> AXIOMSn (A AND B)" using asm by (simp only: AXIOMS_preserved)
       
  5723     }
       
  5724     moreover
       
  5725     { assume "(x):M \<in> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)"
       
  5726       then have "(x):M' \<in> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5727     }
       
  5728     moreover
       
  5729     { assume "(x):M \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)"
       
  5730       then obtain y' N' where eq: "M = AndL1 (y').N' x" 
       
  5731                              and fin: "fin (AndL1 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
       
  5732         by (erule_tac ANDLEFT1_elim, blast)
       
  5733       from eq asm obtain N'' where eq': "M' = AndL1 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
       
  5734         using a_star_redu_AndL1_elim by blast
       
  5735       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       
  5736       moreover
       
  5737       from imp red1 have "(y'):N'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
       
  5738       ultimately have "(x):M' \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)" using eq' by (simp, blast) 
       
  5739     }
       
  5740      moreover
       
  5741     { assume "(x):M \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)"
       
  5742       then obtain y' N' where eq: "M = AndL2 (y').N' x" 
       
  5743                              and fin: "fin (AndL2 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(B)\<parallel>"
       
  5744         by (erule_tac ANDLEFT2_elim, blast)
       
  5745       from eq asm obtain N'' where eq': "M' = AndL2 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
       
  5746         using a_star_redu_AndL2_elim by blast
       
  5747       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       
  5748       moreover
       
  5749       from imp red1 have "(y'):N'' \<in> \<parallel>(B)\<parallel>" using ih4 by simp
       
  5750       ultimately have "(x):M' \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" using eq' by (simp, blast) 
       
  5751     }
       
  5752     ultimately have "(x):M' \<in> AXIOMSn (A AND B) \<union> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)
       
  5753                                \<union> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>) \<union> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" by blast
       
  5754     then have "(x):M' \<in> NEGn (A AND B) (\<parallel><A AND B>\<parallel>)" by simp
       
  5755     then show "(x):M' \<in> (\<parallel>(A AND B)\<parallel>)" using NEG_simp by blast
       
  5756   }
       
  5757 next    
       
  5758  case (OR A B)
       
  5759   have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
       
  5760   have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
       
  5761   have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
       
  5762   have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
       
  5763   { case 1 
       
  5764     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5765     have "<a>:M \<in> \<parallel><A OR B>\<parallel>" by fact
       
  5766     then have "<a>:M \<in> NEGc (A OR B) (\<parallel>(A OR B)\<parallel>)" by simp
       
  5767     then have "<a>:M \<in> AXIOMSc (A OR B) \<union> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>) 
       
  5768                           \<union> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>) \<union> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" by simp
       
  5769     moreover
       
  5770     { assume "<a>:M \<in> AXIOMSc (A OR B)"
       
  5771       then have "<a>:M' \<in> AXIOMSc (A OR B)" using asm by (simp only: AXIOMS_preserved)
       
  5772     }
       
  5773     moreover
       
  5774     { assume "<a>:M \<in> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)"
       
  5775       then have "<a>:M' \<in> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5776     }
       
  5777     moreover
       
  5778     { assume "<a>:M \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)"
       
  5779       then obtain a' N' where eq: "M = OrR1 <a'>.N' a" 
       
  5780                               and fic: "fic (OrR1 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><A>\<parallel>"
       
  5781         using ORRIGHT1_elim by blast
       
  5782       from eq asm obtain N'' where eq': "M' = OrR1 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''" 
       
  5783         using a_star_redu_OrR1_elim by blast
       
  5784       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       
  5785       moreover
       
  5786       from imp1 red1 have "<a'>:N'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
       
  5787       ultimately have "<a>:M' \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)" using eq' by (simp, blast) 
       
  5788     }
       
  5789     moreover
       
  5790     { assume "<a>:M \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)"
       
  5791       then obtain a' N' where eq: "M = OrR2 <a'>.N' a" 
       
  5792                               and fic: "fic (OrR2 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><B>\<parallel>"
       
  5793         using ORRIGHT2_elim by blast
       
  5794       from eq asm obtain N'' where eq': "M' = OrR2 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''" 
       
  5795         using a_star_redu_OrR2_elim by blast
       
  5796       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       
  5797       moreover
       
  5798       from imp1 red1 have "<a'>:N'' \<in> \<parallel><B>\<parallel>" using ih3 by simp
       
  5799       ultimately have "<a>:M' \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" using eq' by (simp, blast) 
       
  5800     }
       
  5801     ultimately have "<a>:M' \<in> AXIOMSc (A OR B) \<union> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)
       
  5802                                 \<union> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>) \<union> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" by blast
       
  5803     then have "<a>:M' \<in> NEGc (A OR B) (\<parallel>(A OR B)\<parallel>)" by simp
       
  5804     then show "<a>:M' \<in> (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
       
  5805   next
       
  5806     case 2
       
  5807     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5808     have "(x):M \<in> \<parallel>(A OR B)\<parallel>" by fact
       
  5809     then have "(x):M \<in> NEGn (A OR B) (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
       
  5810     then have "(x):M \<in> AXIOMSn (A OR B) \<union> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>) 
       
  5811                                      \<union> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" by simp
       
  5812     moreover
       
  5813     { assume "(x):M \<in> AXIOMSn (A OR B)"
       
  5814       then have "(x):M' \<in> AXIOMSn (A OR B)" using asm by (simp only: AXIOMS_preserved)
       
  5815     }
       
  5816     moreover
       
  5817     { assume "(x):M \<in> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)"
       
  5818       then have "(x):M' \<in> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5819     }
       
  5820     moreover
       
  5821     { assume "(x):M \<in> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)"
       
  5822       then obtain y' T' z' N' where eq: "M = OrL (y').T' (z').N' x" 
       
  5823                              and fin: "fin (OrL (y').T' (z').N' x) x" 
       
  5824                              and imp1: "(y'):T' \<in> \<parallel>(A)\<parallel>" and imp2: "(z'):N' \<in> \<parallel>(B)\<parallel>"
       
  5825         by (erule_tac ORLEFT_elim, blast)
       
  5826       from eq asm obtain T'' N'' where eq': "M' = OrL (y').T'' (z').N'' x" 
       
  5827                 and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''"
       
  5828         using a_star_redu_OrL_elim by blast
       
  5829       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       
  5830       moreover
       
  5831       from imp1 red1 have "(y'):T'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
       
  5832       moreover
       
  5833       from imp2 red2 have "(z'):N'' \<in> \<parallel>(B)\<parallel>" using ih4 by simp
       
  5834       ultimately have "(x):M' \<in> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" using eq' by (simp, blast) 
       
  5835     }
       
  5836     ultimately have "(x):M' \<in> AXIOMSn (A OR B) \<union> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)
       
  5837                                \<union> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" by blast
       
  5838     then have "(x):M' \<in> NEGn (A OR B) (\<parallel><A OR B>\<parallel>)" by simp
       
  5839     then show "(x):M' \<in> (\<parallel>(A OR B)\<parallel>)" using NEG_simp by blast
       
  5840   }
       
  5841 next
       
  5842   case (NOT A)
       
  5843   have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
       
  5844   have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
       
  5845   { case 1 
       
  5846     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5847     have "<a>:M \<in> \<parallel><NOT A>\<parallel>" by fact
       
  5848     then have "<a>:M \<in> NEGc (NOT A) (\<parallel>(NOT A)\<parallel>)" by simp
       
  5849     then have "<a>:M \<in> AXIOMSc (NOT A) \<union> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>) 
       
  5850                                               \<union> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)" by simp
       
  5851     moreover
       
  5852     { assume "<a>:M \<in> AXIOMSc (NOT A)"
       
  5853       then have "<a>:M' \<in> AXIOMSc (NOT A)" using asm by (simp only: AXIOMS_preserved)
       
  5854     }
       
  5855     moreover
       
  5856     { assume "<a>:M \<in> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>)"
       
  5857       then have "<a>:M' \<in> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5858     }
       
  5859     moreover
       
  5860     { assume "<a>:M \<in> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)"
       
  5861       then obtain y' N' where eq: "M = NotR (y').N' a" 
       
  5862                               and fic: "fic (NotR (y').N' a) a" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
       
  5863         using NOTRIGHT_elim by blast
       
  5864       from eq asm obtain N'' where eq': "M' = NotR (y').N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''" 
       
  5865         using a_star_redu_NotR_elim by blast
       
  5866       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       
  5867       moreover
       
  5868       from imp red have "(y'):N'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
       
  5869       ultimately have "<a>:M' \<in> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)" using eq' by (simp, blast) 
       
  5870     }
       
  5871     ultimately have "<a>:M' \<in> AXIOMSc (NOT A) \<union> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>)
       
  5872                                               \<union> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)" by blast
       
  5873     then have "<a>:M' \<in> NEGc (NOT A) (\<parallel>(NOT A)\<parallel>)" by simp
       
  5874     then show "<a>:M' \<in> (\<parallel><NOT A>\<parallel>)" using NEG_simp by blast
       
  5875   next
       
  5876     case 2
       
  5877     have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
       
  5878     have "(x):M \<in> \<parallel>(NOT A)\<parallel>" by fact
       
  5879     then have "(x):M \<in> NEGn (NOT A) (\<parallel><NOT A>\<parallel>)" using NEG_simp by blast
       
  5880     then have "(x):M \<in> AXIOMSn (NOT A) \<union> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>) 
       
  5881                                      \<union> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)" by simp
       
  5882     moreover
       
  5883     { assume "(x):M \<in> AXIOMSn (NOT A)"
       
  5884       then have "(x):M' \<in> AXIOMSn (NOT A)" using asm by (simp only: AXIOMS_preserved)
       
  5885     }
       
  5886     moreover
       
  5887     { assume "(x):M \<in> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>)"
       
  5888       then have "(x):M' \<in> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>)" using asm by (simp only: BINDING_preserved)
       
  5889     }
       
  5890     moreover
       
  5891     { assume "(x):M \<in> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)"
       
  5892       then obtain a' N' where eq: "M = NotL <a'>.N' x" 
       
  5893                              and fin: "fin (NotL <a'>.N' x) x" and imp: "<a'>:N' \<in> \<parallel><A>\<parallel>"
       
  5894         by (erule_tac NOTLEFT_elim, blast)
       
  5895       from eq asm obtain N'' where eq': "M' = NotL <a'>.N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
       
  5896         using a_star_redu_NotL_elim by blast
       
  5897       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       
  5898       moreover
       
  5899       from imp red1 have "<a'>:N'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
       
  5900       ultimately have "(x):M' \<in> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)" using eq' by (simp, blast) 
       
  5901     }
       
  5902     ultimately have "(x):M' \<in> AXIOMSn (NOT A) \<union> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>)
       
  5903                                \<union> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)" by blast
       
  5904     then have "(x):M' \<in> NEGn (NOT A) (\<parallel><NOT A>\<parallel>)" by simp
       
  5905     then show "(x):M' \<in> (\<parallel>(NOT A)\<parallel>)" using NEG_simp by blast
       
  5906   }
       
  5907 qed
       
  5908 
       
  5909 lemma CANDs_preserved_single:
       
  5910   shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
       
  5911   and   "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>"
       
  5912 by (auto simp add: a_starI CANDs_preserved)
       
  5913 
       
  5914 lemma fic_CANDS:
       
  5915   assumes a: "\<not>fic M a"
       
  5916   and     b: "<a>:M \<in> \<parallel><B>\<parallel>"
       
  5917   shows "<a>:M \<in> AXIOMSc B \<or> <a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
       
  5918 using a b
       
  5919 apply(nominal_induct B rule: ty.strong_induct)
       
  5920 apply(simp)
       
  5921 apply(simp)
       
  5922 apply(erule disjE)
       
  5923 apply(simp)
       
  5924 apply(erule disjE)
       
  5925 apply(simp)
       
  5926 apply(auto simp add: ctrm.inject)[1]
       
  5927 apply(simp add: alpha)
       
  5928 apply(erule disjE)
       
  5929 apply(simp)
       
  5930 apply(auto simp add: calc_atm)[1]
       
  5931 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
       
  5932 apply(simp add: calc_atm)
       
  5933 apply(simp)
       
  5934 apply(erule disjE)
       
  5935 apply(simp)
       
  5936 apply(erule disjE)
       
  5937 apply(simp)
       
  5938 apply(auto simp add: ctrm.inject)[1]
       
  5939 apply(simp add: alpha)
       
  5940 apply(erule disjE)
       
  5941 apply(simp)
       
  5942 apply(erule conjE)+
       
  5943 apply(simp)
       
  5944 apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
       
  5945 apply(simp add: calc_atm)
       
  5946 apply(simp)
       
  5947 apply(erule disjE)
       
  5948 apply(simp)
       
  5949 apply(erule disjE)
       
  5950 apply(simp)
       
  5951 apply(auto simp add: ctrm.inject)[1]
       
  5952 apply(simp add: alpha)
       
  5953 apply(erule disjE)
       
  5954 apply(simp)
       
  5955 apply(erule conjE)+
       
  5956 apply(simp)
       
  5957 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  5958 apply(simp add: calc_atm)
       
  5959 apply(simp add: alpha)
       
  5960 apply(erule disjE)
       
  5961 apply(simp)
       
  5962 apply(erule conjE)+
       
  5963 apply(simp)
       
  5964 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  5965 apply(simp add: calc_atm)
       
  5966 apply(simp)
       
  5967 apply(erule disjE)
       
  5968 apply(simp)
       
  5969 apply(erule disjE)
       
  5970 apply(simp)
       
  5971 apply(auto simp add: ctrm.inject)[1]
       
  5972 apply(simp add: alpha)
       
  5973 apply(erule disjE)
       
  5974 apply(simp)
       
  5975 apply(erule conjE)+
       
  5976 apply(simp)
       
  5977 apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
       
  5978 apply(simp add: calc_atm)
       
  5979 done
       
  5980 
       
  5981 lemma fin_CANDS_aux:
       
  5982   assumes a: "\<not>fin M x"
       
  5983   and     b: "(x):M \<in> (NEGn B (\<parallel><B>\<parallel>))"
       
  5984   shows "(x):M \<in> AXIOMSn B \<or> (x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
       
  5985 using a b
       
  5986 apply(nominal_induct B rule: ty.strong_induct)
       
  5987 apply(simp)
       
  5988 apply(simp)
       
  5989 apply(erule disjE)
       
  5990 apply(simp)
       
  5991 apply(erule disjE)
       
  5992 apply(simp)
       
  5993 apply(auto simp add: ntrm.inject)[1]
       
  5994 apply(simp add: alpha)
       
  5995 apply(erule disjE)
       
  5996 apply(simp)
       
  5997 apply(auto simp add: calc_atm)[1]
       
  5998 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
       
  5999 apply(simp add: calc_atm)
       
  6000 apply(simp)
       
  6001 apply(erule disjE)
       
  6002 apply(simp)
       
  6003 apply(erule disjE)
       
  6004 apply(simp)
       
  6005 apply(auto simp add: ntrm.inject)[1]
       
  6006 apply(simp add: alpha)
       
  6007 apply(erule disjE)
       
  6008 apply(simp)
       
  6009 apply(erule conjE)+
       
  6010 apply(simp)
       
  6011 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  6012 apply(simp add: calc_atm)
       
  6013 apply(simp add: alpha)
       
  6014 apply(erule disjE)
       
  6015 apply(simp)
       
  6016 apply(erule conjE)+
       
  6017 apply(simp)
       
  6018 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  6019 apply(simp add: calc_atm)
       
  6020 apply(simp)
       
  6021 apply(erule disjE)
       
  6022 apply(simp)
       
  6023 apply(erule disjE)
       
  6024 apply(simp)
       
  6025 apply(auto simp add: ntrm.inject)[1]
       
  6026 apply(simp add: alpha)
       
  6027 apply(erule disjE)
       
  6028 apply(simp)
       
  6029 apply(erule conjE)+
       
  6030 apply(simp)
       
  6031 apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
       
  6032 apply(simp add: calc_atm)
       
  6033 apply(simp)
       
  6034 apply(erule disjE)
       
  6035 apply(simp)
       
  6036 apply(erule disjE)
       
  6037 apply(simp)
       
  6038 apply(auto simp add: ntrm.inject)[1]
       
  6039 apply(simp add: alpha)
       
  6040 apply(erule disjE)
       
  6041 apply(simp)
       
  6042 apply(erule conjE)+
       
  6043 apply(simp)
       
  6044 apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
       
  6045 apply(simp add: calc_atm)
       
  6046 done
       
  6047 
       
  6048 lemma fin_CANDS:
       
  6049   assumes a: "\<not>fin M x"
       
  6050   and     b: "(x):M \<in> (\<parallel>(B)\<parallel>)"
       
  6051   shows "(x):M \<in> AXIOMSn B \<or> (x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
       
  6052 apply(rule fin_CANDS_aux)
       
  6053 apply(rule a)
       
  6054 apply(rule NEG_elim)
       
  6055 apply(rule b)
       
  6056 done
       
  6057 
       
  6058 lemma BINDING_implies_CAND:
       
  6059   shows "<c>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> <c>:M \<in> (\<parallel><B>\<parallel>)"
       
  6060   and   "(x):N \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> (x):N \<in> (\<parallel>(B)\<parallel>)"
       
  6061 apply -
       
  6062 apply(nominal_induct B rule: ty.strong_induct)
       
  6063 apply(auto)
       
  6064 apply(rule NEG_intro)
       
  6065 apply(nominal_induct B rule: ty.strong_induct)
       
  6066 apply(auto)
       
  6067 done
       
  6068 
       
  6069 end