src/HOL/Nominal/Examples/Class3.thy
changeset 72166 bb37571139bf
parent 71989 bad75618fb82
equal deleted inserted replaced
72165:605f151585e0 72166:bb37571139bf
     1 theory Class3
     1 theory Class3
     2 imports Class2
     2   imports Class2
     3 begin
     3 begin
     4 
     4 
     5 text \<open>3rd Main Lemma\<close>
     5 text \<open>3rd Main Lemma\<close>
     6 
     6 
     7 lemma Cut_a_redu_elim:
     7 lemma Cut_a_redu_elim:
     8   assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>a R"
     8   assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>a R"
     9   shows "(\<exists>M'. R = Cut <a>.M' (x).N \<and> M \<longrightarrow>\<^sub>a M') \<or>
     9   shows "(\<exists>M'. R = Cut <a>.M' (x).N \<and> M \<longrightarrow>\<^sub>a M') \<or>
    10          (\<exists>N'. R = Cut <a>.M (x).N' \<and> N \<longrightarrow>\<^sub>a N') \<or>
    10          (\<exists>N'. R = Cut <a>.M (x).N' \<and> N \<longrightarrow>\<^sub>a N') \<or>
    11          (Cut <a>.M (x).N \<longrightarrow>\<^sub>c R) \<or>
    11          (Cut <a>.M (x).N \<longrightarrow>\<^sub>c R) \<or>
    12          (Cut <a>.M (x).N \<longrightarrow>\<^sub>l R)"
    12          (Cut <a>.M (x).N \<longrightarrow>\<^sub>l R)"
    13 using a
    13   using a
    14 apply(erule_tac a_redu.cases)
    14   apply(erule_tac a_redu.cases)
    15 apply(simp_all)
    15                   apply(simp_all)
    16 apply(simp_all add: trm.inject)
    16    apply(simp_all add: trm.inject)
    17 apply(rule disjI1)
    17    apply(rule disjI1)
    18 apply(auto simp add: alpha)[1]
    18    apply(auto simp add: alpha)[1]
    19 apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
    19     apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
    20 apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    20     apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    21 apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
    21    apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
    22 apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    22    apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    23 apply(rule disjI2)
    23   apply(rule disjI2)
    24 apply(rule disjI1)
    24   apply(rule disjI1)
    25 apply(auto simp add: alpha)[1]
    25   apply(auto simp add: alpha)[1]
    26 apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
    26    apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
    27 apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    27    apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    28 apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
    28   apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
    29 apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    29   apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
    30 done
    30   done
    31 
    31 
    32 lemma Cut_c_redu_elim:
    32 lemma Cut_c_redu_elim:
    33   assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>c R"
    33   assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>c R"
    34   shows "(R = M{a:=(x).N} \<and> \<not>fic M a) \<or>
    34   shows "(R = M{a:=(x).N} \<and> \<not>fic M a) \<or>
    35          (R = N{x:=<a>.M} \<and> \<not>fin N x)"
    35          (R = N{x:=<a>.M} \<and> \<not>fin N x)"
    36 using a
    36   using a
    37 apply(erule_tac c_redu.cases)
    37   apply(erule_tac c_redu.cases)
    38 apply(simp_all)
    38    apply(simp_all)
    39 apply(simp_all add: trm.inject)
    39    apply(simp_all add: trm.inject)
    40 apply(rule disjI1)
    40    apply(rule disjI1)
    41 apply(auto simp add: alpha)[1]
    41    apply(auto simp add: alpha)[1]
    42 apply(simp add: subst_rename fresh_atm)
    42        apply(simp add: subst_rename fresh_atm)
    43 apply(simp add: subst_rename fresh_atm)
    43       apply(simp add: subst_rename fresh_atm)
    44 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
    44      apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
    45 apply(perm_simp)
    45      apply(perm_simp)
    46 apply(simp add: subst_rename fresh_atm fresh_prod)
    46     apply(simp add: subst_rename fresh_atm fresh_prod)
    47 apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
    47    apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
    48 apply(perm_simp)
    48    apply(perm_simp)
    49 apply(rule disjI2)
    49   apply(rule disjI2)
    50 apply(auto simp add: alpha)[1]
    50   apply(auto simp add: alpha)[1]
    51 apply(simp add: subst_rename fresh_atm)
    51       apply(simp add: subst_rename fresh_atm)
    52 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
    52      apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
    53 apply(perm_simp)
    53      apply(perm_simp)
    54 apply(simp add: subst_rename fresh_atm fresh_prod)
    54     apply(simp add: subst_rename fresh_atm fresh_prod)
    55 apply(simp add: subst_rename fresh_atm fresh_prod)
    55    apply(simp add: subst_rename fresh_atm fresh_prod)
    56 apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
    56   apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
    57 apply(perm_simp)
    57   apply(perm_simp)
    58 done
    58   done
    59 
    59 
    60 lemma not_fic_crename_aux:
    60 lemma not_fic_crename_aux:
    61   assumes a: "fic M c" "c\<sharp>(a,b)"
    61   assumes a: "fic M c" "c\<sharp>(a,b)"
    62   shows "fic (M[a\<turnstile>c>b]) c" 
    62   shows "fic (M[a\<turnstile>c>b]) c" 
    63 using a
    63   using a
    64 apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
    64   apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
    65 apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
    65              apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
    66 done
    66   done
    67 
    67 
    68 lemma not_fic_crename:
    68 lemma not_fic_crename:
    69   assumes a: "\<not>(fic (M[a\<turnstile>c>b]) c)" "c\<sharp>(a,b)"
    69   assumes a: "\<not>(fic (M[a\<turnstile>c>b]) c)" "c\<sharp>(a,b)"
    70   shows "\<not>(fic M c)" 
    70   shows "\<not>(fic M c)" 
    71 using a
    71   using a
    72 apply(auto dest:  not_fic_crename_aux)
    72   apply(auto dest:  not_fic_crename_aux)
    73 done
    73   done
    74 
    74 
    75 lemma not_fin_crename_aux:
    75 lemma not_fin_crename_aux:
    76   assumes a: "fin M y"
    76   assumes a: "fin M y"
    77   shows "fin (M[a\<turnstile>c>b]) y" 
    77   shows "fin (M[a\<turnstile>c>b]) y" 
    78 using a
    78   using a
    79 apply(nominal_induct M avoiding: a b rule: trm.strong_induct)
    79   apply(nominal_induct M avoiding: a b rule: trm.strong_induct)
    80 apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
    80              apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
    81 done
    81   done
    82 
    82 
    83 lemma not_fin_crename:
    83 lemma not_fin_crename:
    84   assumes a: "\<not>(fin (M[a\<turnstile>c>b]) y)" 
    84   assumes a: "\<not>(fin (M[a\<turnstile>c>b]) y)" 
    85   shows "\<not>(fin M y)" 
    85   shows "\<not>(fin M y)" 
    86 using a
    86   using a
    87 apply(auto dest:  not_fin_crename_aux)
    87   apply(auto dest:  not_fin_crename_aux)
    88 done
    88   done
    89 
    89 
    90 lemma crename_fresh_interesting1:
    90 lemma crename_fresh_interesting1:
    91   fixes c::"coname"
    91   fixes c::"coname"
    92   assumes a: "c\<sharp>(M[a\<turnstile>c>b])" "c\<sharp>(a,b)"
    92   assumes a: "c\<sharp>(M[a\<turnstile>c>b])" "c\<sharp>(a,b)"
    93   shows "c\<sharp>M"
    93   shows "c\<sharp>M"
    94 using a
    94   using a
    95 apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
    95   apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
    96 apply(auto split: if_splits simp add: abs_fresh)
    96              apply(auto split: if_splits simp add: abs_fresh)
    97 done
    97   done
    98 
    98 
    99 lemma crename_fresh_interesting2:
    99 lemma crename_fresh_interesting2:
   100   fixes x::"name"
   100   fixes x::"name"
   101   assumes a: "x\<sharp>(M[a\<turnstile>c>b])" 
   101   assumes a: "x\<sharp>(M[a\<turnstile>c>b])" 
   102   shows "x\<sharp>M"
   102   shows "x\<sharp>M"
   103 using a
   103   using a
   104 apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
   104   apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
   105 apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
   105              apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
   106 done
   106   done
   107 
   107 
   108 
   108 
   109 lemma fic_crename:
   109 lemma fic_crename:
   110   assumes a: "fic (M[a\<turnstile>c>b]) c" "c\<sharp>(a,b)"
   110   assumes a: "fic (M[a\<turnstile>c>b]) c" "c\<sharp>(a,b)"
   111   shows "fic M c" 
   111   shows "fic M c" 
   112 using a
   112   using a
   113 apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
   113   apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
   114 apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
   114              apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
   115            split: if_splits)
   115       split: if_splits)
   116 apply(auto dest: crename_fresh_interesting1 simp add: fresh_prod fresh_atm)
   116        apply(auto dest: crename_fresh_interesting1 simp add: fresh_prod fresh_atm)
   117 done
   117   done
   118 
   118 
   119 lemma fin_crename:
   119 lemma fin_crename:
   120   assumes a: "fin (M[a\<turnstile>c>b]) x"
   120   assumes a: "fin (M[a\<turnstile>c>b]) x"
   121   shows "fin M x" 
   121   shows "fin M x" 
   122 using a
   122   using a
   123 apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
   123   apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
   124 apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
   124              apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
   125            split: if_splits)
   125       split: if_splits)
   126 apply(auto dest: crename_fresh_interesting2 simp add: fresh_prod fresh_atm)
   126         apply(auto dest: crename_fresh_interesting2 simp add: fresh_prod fresh_atm)
   127 done
   127   done
   128 
   128 
   129 lemma crename_Cut:
   129 lemma crename_Cut:
   130   assumes a: "R[a\<turnstile>c>b] = Cut <c>.M (x).N" "c\<sharp>(a,b,N,R)" "x\<sharp>(M,R)"
   130   assumes a: "R[a\<turnstile>c>b] = Cut <c>.M (x).N" "c\<sharp>(a,b,N,R)" "x\<sharp>(M,R)"
   131   shows "\<exists>M' N'. R = Cut <c>.M' (x).N' \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> x\<sharp>M'"
   131   shows "\<exists>M' N'. R = Cut <c>.M' (x).N' \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> x\<sharp>M'"
   132 using a
   132   using a
   133 apply(nominal_induct R avoiding: a b c x M N rule: trm.strong_induct)
   133   apply(nominal_induct R avoiding: a b c x M N rule: trm.strong_induct)
   134 apply(auto split: if_splits)
   134              apply(auto split: if_splits)
   135 apply(simp add: trm.inject)
   135   apply(simp add: trm.inject)
   136 apply(auto simp add: alpha)
   136   apply(auto simp add: alpha)
   137 apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
   137     apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
   138 apply(perm_simp)
   138     apply(perm_simp)
   139 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   139     apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   140 apply(drule sym)
   140     apply(drule sym)
   141 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   141     apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   142 apply(simp add: eqvts calc_atm)
   142     apply(simp add: eqvts calc_atm)
   143 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   143    apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   144 apply(perm_simp)
   144    apply(perm_simp)
   145 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   145    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   146 apply(drule sym)
   146    apply(drule sym)
   147 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   147    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   148 apply(simp add: eqvts calc_atm)
   148    apply(simp add: eqvts calc_atm)
   149 apply(auto simp add: fresh_atm)[1]
   149    apply(auto simp add: fresh_atm)[1]
   150 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   150   apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   151 apply(perm_simp)
   151   apply(perm_simp)
   152 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   152   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   153 apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
   153   apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
   154 apply(perm_simp)
   154   apply(perm_simp)
   155 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   155   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   156 apply(drule sym)
   156   apply(drule sym)
   157 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   157   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   158 apply(simp add: eqvts calc_atm)
   158   apply(simp add: eqvts calc_atm)
   159 apply(auto simp add: fresh_atm)[1]
   159   apply(auto simp add: fresh_atm)[1]
   160 apply(drule sym)
   160    apply(drule sym)
   161 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   161    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   162 apply(simp add: eqvts calc_atm)
   162    apply(simp add: eqvts calc_atm)
   163 apply(drule sym)
   163   apply(drule sym)
   164 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   164   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   165 apply(simp add: eqvts calc_atm)
   165   apply(simp add: eqvts calc_atm)
   166 done
   166   done
   167 
   167 
   168 lemma crename_NotR:
   168 lemma crename_NotR:
   169   assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>(a,b)"
   169   assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>(a,b)"
   170   shows "\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N" 
   170   shows "\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N" 
   171 using a
   171   using a
   172 apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
   172   apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
   173 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   173              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   174 apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   174   apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   175 apply(perm_simp)
   175   apply(perm_simp)
   176 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   176   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   177 apply(drule sym)
   177   apply(drule sym)
   178 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   178   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   179 apply(simp add: eqvts calc_atm)
   179   apply(simp add: eqvts calc_atm)
   180 done
   180   done
   181 
   181 
   182 lemma crename_NotR':
   182 lemma crename_NotR':
   183   assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>a"
   183   assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>a"
   184   shows "(\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N) \<or> (\<exists>N'. (R = NotR (x).N' a) \<and> b=c \<and> N'[a\<turnstile>c>b] = N)"
   184   shows "(\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N) \<or> (\<exists>N'. (R = NotR (x).N' a) \<and> b=c \<and> N'[a\<turnstile>c>b] = N)"
   185 using a
   185   using a
   186 apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
   186   apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
   187 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
   187              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
   188 apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   188    apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   189 apply(perm_simp)
   189    apply(perm_simp)
   190 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   190    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   191 apply(drule sym)
   191    apply(drule sym)
   192 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   192    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   193 apply(simp add: eqvts calc_atm)
   193    apply(simp add: eqvts calc_atm)
   194 apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   194   apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   195 apply(perm_simp)
   195   apply(perm_simp)
   196 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   196   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   197 apply(drule sym)
   197   apply(drule sym)
   198 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   198   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   199 apply(simp add: eqvts calc_atm)
   199   apply(simp add: eqvts calc_atm)
   200 done
   200   done
   201 
   201 
   202 lemma crename_NotR_aux:
   202 lemma crename_NotR_aux:
   203   assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" 
   203   assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" 
   204   shows "(a=c \<and> a=b) \<or> (a\<noteq>c)" 
   204   shows "(a=c \<and> a=b) \<or> (a\<noteq>c)" 
   205 using a
   205   using a
   206 apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
   206   apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
   207 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   207              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   208 done
   208   done
   209 
   209 
   210 lemma crename_NotL:
   210 lemma crename_NotL:
   211   assumes a: "R[a\<turnstile>c>b] = NotL <c>.N y" "c\<sharp>(R,a,b)"
   211   assumes a: "R[a\<turnstile>c>b] = NotL <c>.N y" "c\<sharp>(R,a,b)"
   212   shows "\<exists>N'. (R = NotL <c>.N' y) \<and> N'[a\<turnstile>c>b] = N" 
   212   shows "\<exists>N'. (R = NotL <c>.N' y) \<and> N'[a\<turnstile>c>b] = N" 
   213 using a
   213   using a
   214 apply(nominal_induct R avoiding: a b c y N rule: trm.strong_induct)
   214   apply(nominal_induct R avoiding: a b c y N rule: trm.strong_induct)
   215 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   215              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   216 apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
   216   apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
   217 apply(perm_simp)
   217   apply(perm_simp)
   218 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   218   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   219 apply(drule sym)
   219   apply(drule sym)
   220 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   220   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   221 apply(simp add: eqvts calc_atm)
   221   apply(simp add: eqvts calc_atm)
   222 done
   222   done
   223 
   223 
   224 lemma crename_AndL1:
   224 lemma crename_AndL1:
   225   assumes a: "R[a\<turnstile>c>b] = AndL1 (x).N y" "x\<sharp>R"
   225   assumes a: "R[a\<turnstile>c>b] = AndL1 (x).N y" "x\<sharp>R"
   226   shows "\<exists>N'. (R = AndL1 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
   226   shows "\<exists>N'. (R = AndL1 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
   227 using a
   227   using a
   228 apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
   228   apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
   229 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   229              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   230 apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
   230   apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
   231 apply(perm_simp)
   231   apply(perm_simp)
   232 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   232   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   233 apply(drule sym)
   233   apply(drule sym)
   234 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   234   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   235 apply(simp add: eqvts calc_atm)
   235   apply(simp add: eqvts calc_atm)
   236 done
   236   done
   237 
   237 
   238 lemma crename_AndL2:
   238 lemma crename_AndL2:
   239   assumes a: "R[a\<turnstile>c>b] = AndL2 (x).N y" "x\<sharp>R"
   239   assumes a: "R[a\<turnstile>c>b] = AndL2 (x).N y" "x\<sharp>R"
   240   shows "\<exists>N'. (R = AndL2 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
   240   shows "\<exists>N'. (R = AndL2 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
   241 using a
   241   using a
   242 apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
   242   apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
   243 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   243              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   244 apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
   244   apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
   245 apply(perm_simp)
   245   apply(perm_simp)
   246 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   246   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   247 apply(drule sym)
   247   apply(drule sym)
   248 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   248   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   249 apply(simp add: eqvts calc_atm)
   249   apply(simp add: eqvts calc_atm)
   250 done
   250   done
   251 
   251 
   252 lemma crename_AndR_aux:
   252 lemma crename_AndR_aux:
   253   assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" 
   253   assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" 
   254   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   254   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   255 using a
   255   using a
   256 apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
   256   apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
   257 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   257              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   258 done
   258   done
   259 
   259 
   260 lemma crename_AndR:
   260 lemma crename_AndR:
   261   assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>(a,b)"
   261   assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>(a,b)"
   262   shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
   262   shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
   263 using a
   263   using a
   264 apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
   264   apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
   265 apply(auto split: if_splits simp add: trm.inject alpha)
   265              apply(auto split: if_splits simp add: trm.inject alpha)
   266 apply(simp add: fresh_atm fresh_prod)
   266         apply(simp add: fresh_atm fresh_prod)
   267 apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   267        apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   268 apply(perm_simp)
   268        apply(perm_simp)
   269 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   269        apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   270 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   270       apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   271 apply(perm_simp)
   271       apply(perm_simp)
   272 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   272       apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   273 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   273      apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   274 apply(perm_simp)
   274      apply(perm_simp)
   275 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   275      apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   276 apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   276     apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   277 apply(perm_simp)
   277     apply(perm_simp)
   278 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   278     apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   279 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   279    apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   280 apply(perm_simp)
   280    apply(perm_simp)
   281 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   281    apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   282 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   282   apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   283 apply(perm_simp)
   283   apply(perm_simp)
   284 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   284   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   285 apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   285   apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   286 apply(perm_simp)
   286   apply(perm_simp)
   287 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   287   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   288 apply(drule sym)
   288    apply(drule sym)
   289 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   289    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   290 apply(simp add: eqvts calc_atm)
   290    apply(simp add: eqvts calc_atm)
   291 apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   291   apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   292 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   292   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   293 apply(simp add: eqvts calc_atm)
   293   apply(simp add: eqvts calc_atm)
   294 done
   294   done
   295 
   295 
   296 lemma crename_AndR':
   296 lemma crename_AndR':
   297   assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>a"
   297   assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>a"
   298   shows "(\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M') \<or>
   298   shows "(\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M') \<or>
   299          (\<exists>M' N'. R = AndR <c>.M' <d>.N' a \<and> b=e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M')"
   299          (\<exists>M' N'. R = AndR <c>.M' <d>.N' a \<and> b=e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M')"
   300 using a [[simproc del: defined_all]]
   300   using a [[simproc del: defined_all]]
   301 apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
   301   apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
   302 apply(auto split: if_splits simp add: trm.inject alpha)[1]
   302              apply(auto split: if_splits simp add: trm.inject alpha)[1]
   303 apply(auto split: if_splits simp add: trm.inject alpha)[1]
   303             apply(auto split: if_splits simp add: trm.inject alpha)[1]
   304 apply(auto split: if_splits simp add: trm.inject alpha)[1]
   304            apply(auto split: if_splits simp add: trm.inject alpha)[1]
   305 apply(auto split: if_splits simp add: trm.inject alpha)[1]
   305           apply(auto split: if_splits simp add: trm.inject alpha)[1]
   306 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha)[1]
   306          apply(simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha)[1]
   307 apply(case_tac "coname3=a")
   307          apply(case_tac "coname3=a")
   308 apply(simp)
   308           apply(simp)
   309 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   309           apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   310 apply(perm_simp)
   310           apply(perm_simp)
   311 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   311           apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   312 apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   312           apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   313 apply(perm_simp)
   313           apply(perm_simp)
   314 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
   314           apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
   315 apply(drule sym)
   315            apply(drule sym)
   316 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   316            apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   317 apply(simp add: eqvts calc_atm)
   317            apply(simp add: eqvts calc_atm)
   318 apply(drule_tac s="trm2[a\<turnstile>c>e]" in  sym)
   318           apply(drule_tac s="trm2[a\<turnstile>c>e]" in  sym)
   319 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   319           apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   320 apply(simp add: eqvts calc_atm)
   320           apply(simp add: eqvts calc_atm)
   321 apply(simp)
   321          apply(simp)
   322 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   322          apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
   323 apply(perm_simp)
   323          apply(perm_simp)
   324 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   324          apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   325 apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   325          apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
   326 apply(perm_simp)
   326          apply(perm_simp)
   327 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
   327          apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
   328 apply(drule sym)
   328           apply(drule sym)
   329 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   329           apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   330 apply(simp add: eqvts calc_atm)
   330           apply(simp add: eqvts calc_atm)
   331 apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   331          apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   332 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   332          apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   333 apply(simp add: eqvts calc_atm)
   333          apply(simp add: eqvts calc_atm)
   334 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   334         apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   335 done
   335   done
   336 
   336 
   337 lemma crename_OrR1_aux:
   337 lemma crename_OrR1_aux:
   338   assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.M e" 
   338   assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.M e" 
   339   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   339   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   340 using a
   340   using a
   341 apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
   341   apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
   342 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   342              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   343 done
   343   done
   344 
   344 
   345 lemma crename_OrR1:
   345 lemma crename_OrR1:
   346   assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
   346   assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
   347   shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
   347   shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
   348 using a
   348   using a
   349 apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   349   apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   350 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   350              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   351 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   351   apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   352 apply(perm_simp)
   352   apply(perm_simp)
   353 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   353   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   354 apply(drule sym)
   354   apply(drule sym)
   355 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   355   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   356 apply(simp add: eqvts calc_atm)
   356   apply(simp add: eqvts calc_atm)
   357 done
   357   done
   358 
   358 
   359 lemma crename_OrR1':
   359 lemma crename_OrR1':
   360   assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
   360   assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
   361   shows "(\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
   361   shows "(\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
   362          (\<exists>N'. (R = OrR1 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
   362          (\<exists>N'. (R = OrR1 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
   363 using a
   363   using a
   364 apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   364   apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   365 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   365              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   366 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   366    apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   367 apply(perm_simp)
   367    apply(perm_simp)
   368 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   368    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   369 apply(drule sym)
   369    apply(drule sym)
   370 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   370    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   371 apply(simp add: eqvts calc_atm)
   371    apply(simp add: eqvts calc_atm)
   372 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   372   apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   373 apply(perm_simp)
   373   apply(perm_simp)
   374 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   374   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   375 apply(drule sym)
   375   apply(drule sym)
   376 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   376   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   377 apply(simp add: eqvts calc_atm)
   377   apply(simp add: eqvts calc_atm)
   378 done
   378   done
   379 
   379 
   380 lemma crename_OrR2_aux:
   380 lemma crename_OrR2_aux:
   381   assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.M e" 
   381   assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.M e" 
   382   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   382   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   383 using a
   383   using a
   384 apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
   384   apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
   385 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   385              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   386 done
   386   done
   387 
   387 
   388 lemma crename_OrR2:
   388 lemma crename_OrR2:
   389   assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
   389   assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
   390   shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
   390   shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
   391 using a
   391   using a
   392 apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   392   apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   393 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   393              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   394 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   394   apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   395 apply(perm_simp)
   395   apply(perm_simp)
   396 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   396   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   397 apply(drule sym)
   397   apply(drule sym)
   398 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   398   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   399 apply(simp add: eqvts calc_atm)
   399   apply(simp add: eqvts calc_atm)
   400 done
   400   done
   401 
   401 
   402 lemma crename_OrR2':
   402 lemma crename_OrR2':
   403   assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
   403   assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
   404   shows "(\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
   404   shows "(\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
   405          (\<exists>N'. (R = OrR2 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
   405          (\<exists>N'. (R = OrR2 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
   406 using a
   406   using a
   407 apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   407   apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
   408 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   408              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   409 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   409    apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   410 apply(perm_simp)
   410    apply(perm_simp)
   411 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   411    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   412 apply(drule sym)
   412    apply(drule sym)
   413 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   413    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   414 apply(simp add: eqvts calc_atm)
   414    apply(simp add: eqvts calc_atm)
   415 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   415   apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
   416 apply(perm_simp)
   416   apply(perm_simp)
   417 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   417   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   418 apply(drule sym)
   418   apply(drule sym)
   419 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   419   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   420 apply(simp add: eqvts calc_atm)
   420   apply(simp add: eqvts calc_atm)
   421 done
   421   done
   422 
   422 
   423 lemma crename_OrL:
   423 lemma crename_OrL:
   424   assumes a: "R[a\<turnstile>c>b] = OrL (x).M (y).N z" "x\<sharp>(y,z,N,R)" "y\<sharp>(x,z,M,R)"
   424   assumes a: "R[a\<turnstile>c>b] = OrL (x).M (y).N z" "x\<sharp>(y,z,N,R)" "y\<sharp>(x,z,M,R)"
   425   shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
   425   shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
   426 using a
   426   using a
   427 apply(nominal_induct R avoiding: a b x y z M N rule: trm.strong_induct)
   427   apply(nominal_induct R avoiding: a b x y z M N rule: trm.strong_induct)
   428 apply(auto split: if_splits simp add: trm.inject alpha)
   428              apply(auto split: if_splits simp add: trm.inject alpha)
   429 apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
   429     apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
   430 apply(perm_simp)
   430     apply(perm_simp)
   431 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   431     apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   432 apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
   432    apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
   433 apply(perm_simp)
   433    apply(perm_simp)
   434 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   434    apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   435 apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
   435   apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
   436 apply(perm_simp)
   436   apply(perm_simp)
   437 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   437   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   438 apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
   438   apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
   439 apply(perm_simp)
   439   apply(perm_simp)
   440 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   440   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   441 apply(drule sym)
   441    apply(drule sym)
   442 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   442    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   443 apply(simp add: eqvts calc_atm)
   443    apply(simp add: eqvts calc_atm)
   444 apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   444   apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   445 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   445   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   446 apply(simp add: eqvts calc_atm)
   446   apply(simp add: eqvts calc_atm)
   447 done
   447   done
   448 
   448 
   449 lemma crename_ImpL:
   449 lemma crename_ImpL:
   450   assumes a: "R[a\<turnstile>c>b] = ImpL <c>.M (y).N z" "c\<sharp>(a,b,N,R)" "y\<sharp>(z,M,R)"
   450   assumes a: "R[a\<turnstile>c>b] = ImpL <c>.M (y).N z" "c\<sharp>(a,b,N,R)" "y\<sharp>(z,M,R)"
   451   shows "\<exists>M' N'. R = ImpL <c>.M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> y\<sharp>M'"
   451   shows "\<exists>M' N'. R = ImpL <c>.M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> y\<sharp>M'"
   452 using a
   452   using a
   453 apply(nominal_induct R avoiding: a b c y z M N rule: trm.strong_induct)
   453   apply(nominal_induct R avoiding: a b c y z M N rule: trm.strong_induct)
   454 apply(auto split: if_splits simp add: trm.inject alpha)
   454              apply(auto split: if_splits simp add: trm.inject alpha)
   455 apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
   455     apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
   456 apply(perm_simp)
   456     apply(perm_simp)
   457 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   457     apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   458 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   458    apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   459 apply(perm_simp)
   459    apply(perm_simp)
   460 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   460    apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   461 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   461   apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
   462 apply(perm_simp)
   462   apply(perm_simp)
   463 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   463   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   464 apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
   464   apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
   465 apply(perm_simp)
   465   apply(perm_simp)
   466 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   466   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
   467 apply(drule sym)
   467    apply(drule sym)
   468 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   468    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   469 apply(simp add: eqvts calc_atm)
   469    apply(simp add: eqvts calc_atm)
   470 apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   470   apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
   471 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   471   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   472 apply(simp add: eqvts calc_atm)
   472   apply(simp add: eqvts calc_atm)
   473 done
   473   done
   474 
   474 
   475 lemma crename_ImpR_aux:
   475 lemma crename_ImpR_aux:
   476   assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.M e" 
   476   assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.M e" 
   477   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   477   shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
   478 using a
   478   using a
   479 apply(nominal_induct R avoiding: x a b c e M rule: trm.strong_induct)
   479   apply(nominal_induct R avoiding: x a b c e M rule: trm.strong_induct)
   480 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   480              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
   481 done
   481   done
   482 
   482 
   483 lemma crename_ImpR:
   483 lemma crename_ImpR:
   484   assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)" "x\<sharp>R" 
   484   assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)" "x\<sharp>R" 
   485   shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
   485   shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
   486 using a
   486   using a
   487 apply(nominal_induct R avoiding: a b x c d N rule: trm.strong_induct)
   487   apply(nominal_induct R avoiding: a b x c d N rule: trm.strong_induct)
   488 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
   488              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
   489 apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   489    apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
   490 apply(perm_simp)
   490    apply(perm_simp)
   491 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   491   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
   492 apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
   492   apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
   493 apply(perm_simp)
   493   apply(perm_simp)
   494 apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
   494   apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
   495 apply(drule sym)
   495   apply(drule sym)
   496 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   496   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   497 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   497   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   498 apply(simp add: eqvts calc_atm)
   498   apply(simp add: eqvts calc_atm)
   499 done
   499   done
   500 
   500 
   501 lemma crename_ImpR':
   501 lemma crename_ImpR':
   502   assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "x\<sharp>R" "d\<sharp>a"
   502   assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "x\<sharp>R" "d\<sharp>a"
   503   shows "(\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
   503   shows "(\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
   504          (\<exists>N'. (R = ImpR (x).<c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
   504          (\<exists>N'. (R = ImpR (x).<c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
   505 using a
   505   using a
   506 apply(nominal_induct R avoiding: x a b c d N rule: trm.strong_induct)
   506   apply(nominal_induct R avoiding: x a b c d N rule: trm.strong_induct)
   507 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject abs_perm calc_atm)
   507              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject abs_perm calc_atm)
   508 apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
   508    apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
   509 apply(perm_simp)
   509    apply(perm_simp)
   510 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
   510    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
   511 apply(drule sym)
   511    apply(drule sym)
   512 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   512    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   513 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   513    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   514 apply(simp add: eqvts calc_atm)
   514    apply(simp add: eqvts calc_atm)
   515 apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
   515   apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
   516 apply(perm_simp)
   516   apply(perm_simp)
   517 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
   517   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
   518 apply(drule sym)
   518   apply(drule sym)
   519 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   519   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
   520 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   520   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
   521 apply(simp add: eqvts calc_atm)
   521   apply(simp add: eqvts calc_atm)
   522 done
   522   done
   523 
   523 
   524 lemma crename_ax2:
   524 lemma crename_ax2:
   525   assumes a: "N[a\<turnstile>c>b] = Ax x c"
   525   assumes a: "N[a\<turnstile>c>b] = Ax x c"
   526   shows "\<exists>d. N = Ax x d"
   526   shows "\<exists>d. N = Ax x d"
   527 using a
   527   using a
   528 apply(nominal_induct N avoiding: a b rule: trm.strong_induct)
   528   apply(nominal_induct N avoiding: a b rule: trm.strong_induct)
   529 apply(auto split: if_splits)
   529              apply(auto split: if_splits)
   530 apply(simp add: trm.inject)
   530   apply(simp add: trm.inject)
   531 done
   531   done
   532 
   532 
   533 lemma crename_interesting1:
   533 lemma crename_interesting1:
   534   assumes a: "distinct [a,b,c]"
   534   assumes a: "distinct [a,b,c]"
   535   shows "M[a\<turnstile>c>c][c\<turnstile>c>b] = M[c\<turnstile>c>b][a\<turnstile>c>b]"
   535   shows "M[a\<turnstile>c>c][c\<turnstile>c>b] = M[c\<turnstile>c>b][a\<turnstile>c>b]"
   536 using a
   536   using a
   537 apply(nominal_induct M avoiding: a c b rule: trm.strong_induct)
   537   apply(nominal_induct M avoiding: a c b rule: trm.strong_induct)
   538 apply(auto simp add: rename_fresh simp add: trm.inject alpha)
   538              apply(auto simp add: rename_fresh simp add: trm.inject alpha)
   539 apply(blast)
   539      apply(blast)
   540 apply(rotate_tac 12)
   540     apply(rotate_tac 12)
   541 apply(drule_tac x="a" in meta_spec)
   541     apply(drule_tac x="a" in meta_spec)
   542 apply(rotate_tac 15)
   542     apply(rotate_tac 15)
   543 apply(drule_tac x="c" in meta_spec)
   543     apply(drule_tac x="c" in meta_spec)
   544 apply(rotate_tac 15)
   544     apply(rotate_tac 15)
   545 apply(drule_tac x="b" in meta_spec)
   545     apply(drule_tac x="b" in meta_spec)
   546 apply(blast)
   546     apply(blast)
   547 apply(blast)
   547    apply(blast)
   548 apply(blast)
   548   apply(blast)
   549 done
   549   done
   550 
   550 
   551 lemma crename_interesting2:
   551 lemma crename_interesting2:
   552   assumes a: "a\<noteq>c" "a\<noteq>d" "a\<noteq>b" "c\<noteq>d" "b\<noteq>c"
   552   assumes a: "a\<noteq>c" "a\<noteq>d" "a\<noteq>b" "c\<noteq>d" "b\<noteq>c"
   553   shows "M[a\<turnstile>c>b][c\<turnstile>c>d] = M[c\<turnstile>c>d][a\<turnstile>c>b]"
   553   shows "M[a\<turnstile>c>b][c\<turnstile>c>d] = M[c\<turnstile>c>d][a\<turnstile>c>b]"
   554 using a
   554   using a
   555 apply(nominal_induct M avoiding: a c b d rule: trm.strong_induct)
   555   apply(nominal_induct M avoiding: a c b d rule: trm.strong_induct)
   556 apply(auto simp add: rename_fresh simp add: trm.inject alpha)
   556              apply(auto simp add: rename_fresh simp add: trm.inject alpha)
   557 done
   557   done
   558 
   558 
   559 lemma crename_interesting3:
   559 lemma crename_interesting3:
   560   shows "M[a\<turnstile>c>c][x\<turnstile>n>y] = M[x\<turnstile>n>y][a\<turnstile>c>c]"
   560   shows "M[a\<turnstile>c>c][x\<turnstile>n>y] = M[x\<turnstile>n>y][a\<turnstile>c>c]"
   561 apply(nominal_induct M avoiding: a c x y rule: trm.strong_induct)
   561   apply(nominal_induct M avoiding: a c x y rule: trm.strong_induct)
   562 apply(auto simp add: rename_fresh simp add: trm.inject alpha)
   562              apply(auto simp add: rename_fresh simp add: trm.inject alpha)
   563 done
   563   done
   564 
   564 
   565 lemma crename_credu:
   565 lemma crename_credu:
   566   assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>c M'"
   566   assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>c M'"
   567   shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>c M0"
   567   shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>c M0"
   568 using a
   568   using a
   569 apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: c_redu.strong_induct)
   569   apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: c_redu.strong_induct)
   570 apply(drule sym)
   570    apply(drule sym)
   571 apply(drule crename_Cut)
   571    apply(drule crename_Cut)
   572 apply(simp)
   572      apply(simp)
   573 apply(simp)
   573     apply(simp)
   574 apply(auto)
   574    apply(auto)
   575 apply(rule_tac x="M'{a:=(x).N'}" in exI)
   575    apply(rule_tac x="M'{a:=(x).N'}" in exI)
   576 apply(rule conjI)
   576    apply(rule conjI)
   577 apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
   577     apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
   578 apply(rule c_redu.intros)
   578    apply(rule c_redu.intros)
   579 apply(auto dest: not_fic_crename)[1]
   579      apply(auto dest: not_fic_crename)[1]
   580 apply(simp add: abs_fresh)
   580     apply(simp add: abs_fresh)
   581 apply(simp add: abs_fresh)
   581    apply(simp add: abs_fresh)
   582 apply(drule sym)
   582   apply(drule sym)
   583 apply(drule crename_Cut)
   583   apply(drule crename_Cut)
   584 apply(simp)
   584     apply(simp)
   585 apply(simp)
   585    apply(simp)
   586 apply(auto)
   586   apply(auto)
   587 apply(rule_tac x="N'{x:=<a>.M'}" in exI)
   587   apply(rule_tac x="N'{x:=<a>.M'}" in exI)
   588 apply(rule conjI)
   588   apply(rule conjI)
   589 apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
   589    apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
   590 apply(rule c_redu.intros)
   590   apply(rule c_redu.intros)
   591 apply(auto dest: not_fin_crename)[1]
   591     apply(auto dest: not_fin_crename)[1]
   592 apply(simp add: abs_fresh)
   592    apply(simp add: abs_fresh)
   593 apply(simp add: abs_fresh)
   593   apply(simp add: abs_fresh)
   594 done
   594   done
   595 
   595 
   596 lemma crename_lredu:
   596 lemma crename_lredu:
   597   assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>l M'"
   597   assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>l M'"
   598   shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>l M0"
   598   shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>l M0"
   599 using a
   599   using a
   600 apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: l_redu.strong_induct)
   600   apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: l_redu.strong_induct)
   601 apply(drule sym)
   601          apply(drule sym)
   602 apply(drule crename_Cut)
   602          apply(drule crename_Cut)
   603 apply(simp add: fresh_prod fresh_atm)
   603            apply(simp add: fresh_prod fresh_atm)
   604 apply(simp)
   604           apply(simp)
   605 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   605          apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   606 apply(case_tac "aa=ba")
   606          apply(case_tac "aa=ba")
   607 apply(simp add: crename_id)
   607           apply(simp add: crename_id)
   608 apply(rule l_redu.intros)
   608           apply(rule l_redu.intros)
   609 apply(simp)
   609             apply(simp)
   610 apply(simp add: fresh_atm)
   610            apply(simp add: fresh_atm)
   611 apply(assumption)
   611           apply(assumption)
   612 apply(frule crename_ax2)
   612          apply(frule crename_ax2)
   613 apply(auto)[1]
   613          apply(auto)[1]
   614 apply(case_tac "d=aa")
   614          apply(case_tac "d=aa")
   615 apply(simp add: trm.inject)
   615           apply(simp add: trm.inject)
   616 apply(rule_tac x="M'[a\<turnstile>c>aa]" in exI)
   616           apply(rule_tac x="M'[a\<turnstile>c>aa]" in exI)
   617 apply(rule conjI)
   617           apply(rule conjI)
   618 apply(rule crename_interesting1)
   618            apply(rule crename_interesting1)
   619 apply(simp)
   619            apply(simp)
   620 apply(rule l_redu.intros)
   620           apply(rule l_redu.intros)
   621 apply(simp)
   621             apply(simp)
   622 apply(simp add: fresh_atm)
   622            apply(simp add: fresh_atm)
   623 apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
   623           apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
   624 apply(simp add: trm.inject)
   624          apply(simp add: trm.inject)
   625 apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
   625          apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
   626 apply(rule conjI)
   626          apply(rule conjI)
   627 apply(rule crename_interesting2)
   627           apply(rule crename_interesting2)
   628 apply(simp)
   628               apply(simp)
   629 apply(simp)
   629              apply(simp)
   630 apply(simp)
   630             apply(simp)
   631 apply(simp)
   631            apply(simp)
   632 apply(simp)
   632           apply(simp)
   633 apply(rule l_redu.intros)
   633          apply(rule l_redu.intros)
   634 apply(simp)
   634            apply(simp)
   635 apply(simp add: fresh_atm)
   635           apply(simp add: fresh_atm)
   636 apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
   636          apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
   637 apply(drule sym)
   637         apply(drule sym)
   638 apply(drule crename_Cut)
   638         apply(drule crename_Cut)
   639 apply(simp add: fresh_prod fresh_atm)
   639           apply(simp add: fresh_prod fresh_atm)
   640 apply(simp add: fresh_prod fresh_atm)
   640          apply(simp add: fresh_prod fresh_atm)
   641 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   641         apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   642 apply(case_tac "aa=b")
   642         apply(case_tac "aa=b")
   643 apply(simp add: crename_id)
   643          apply(simp add: crename_id)
   644 apply(rule l_redu.intros)
   644          apply(rule l_redu.intros)
   645 apply(simp)
   645            apply(simp)
   646 apply(simp add: fresh_atm)
   646           apply(simp add: fresh_atm)
   647 apply(assumption)
   647          apply(assumption)
   648 apply(frule crename_ax2)
   648         apply(frule crename_ax2)
   649 apply(auto)[1]
   649         apply(auto)[1]
   650 apply(case_tac "d=aa")
   650         apply(case_tac "d=aa")
   651 apply(simp add: trm.inject)
   651          apply(simp add: trm.inject)
   652 apply(simp add: trm.inject)
   652         apply(simp add: trm.inject)
   653 apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
   653         apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
   654 apply(rule conjI)
   654         apply(rule conjI)
   655 apply(rule sym)
   655          apply(rule sym)
   656 apply(rule crename_interesting3)
   656          apply(rule crename_interesting3)
   657 apply(rule l_redu.intros)
   657         apply(rule l_redu.intros)
   658 apply(simp)
   658           apply(simp)
   659 apply(simp add: fresh_atm)
   659          apply(simp add: fresh_atm)
   660 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   660         apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   661 (* LNot *)
   661     (* LNot *)
   662 apply(drule sym)
   662        apply(drule sym)
   663 apply(drule crename_Cut)
   663        apply(drule crename_Cut)
   664 apply(simp add: fresh_prod abs_fresh fresh_atm)
   664          apply(simp add: fresh_prod abs_fresh fresh_atm)
   665 apply(simp add: fresh_prod abs_fresh fresh_atm)
   665         apply(simp add: fresh_prod abs_fresh fresh_atm)
   666 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   666        apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   667 apply(drule crename_NotR)
   667        apply(drule crename_NotR)
   668 apply(simp add: fresh_prod abs_fresh fresh_atm)
   668          apply(simp add: fresh_prod abs_fresh fresh_atm)
   669 apply(simp add: fresh_prod abs_fresh fresh_atm)
   669         apply(simp add: fresh_prod abs_fresh fresh_atm)
   670 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   670        apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   671 apply(drule crename_NotL)
   671        apply(drule crename_NotL)
   672 apply(simp add: fresh_prod abs_fresh fresh_atm)
   672         apply(simp add: fresh_prod abs_fresh fresh_atm)
   673 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   673        apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   674 apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
   674        apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
   675 apply(simp add: fresh_atm)[1]
   675        apply(simp add: fresh_atm)[1]
   676 apply(rule l_redu.intros)
   676        apply(rule l_redu.intros)
   677 apply(auto simp add: fresh_prod intro: crename_fresh_interesting2)[1]
   677             apply(auto simp add: fresh_prod intro: crename_fresh_interesting2)[1]
   678 apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   678            apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   679 apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   679           apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   680 apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   680          apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   681 apply(simp add: fresh_atm)
   681         apply(simp add: fresh_atm)
   682 apply(simp add: fresh_atm)
   682        apply(simp add: fresh_atm)
   683 (* LAnd1 *)
   683     (* LAnd1 *)
   684 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   684       apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   685 apply(drule sym)
   685       apply(drule sym)
   686 apply(drule crename_Cut)
   686       apply(drule crename_Cut)
   687 apply(simp add: fresh_prod abs_fresh fresh_atm)
   687         apply(simp add: fresh_prod abs_fresh fresh_atm)
   688 apply(simp add: fresh_prod abs_fresh fresh_atm)
   688        apply(simp add: fresh_prod abs_fresh fresh_atm)
   689 apply(auto)[1]
   689       apply(auto)[1]
   690 apply(drule crename_AndR)
   690       apply(drule crename_AndR)
   691 apply(simp add: fresh_prod abs_fresh fresh_atm)
   691          apply(simp add: fresh_prod abs_fresh fresh_atm)
   692 apply(simp add: fresh_prod abs_fresh fresh_atm)
   692         apply(simp add: fresh_prod abs_fresh fresh_atm)
   693 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   693        apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   694 apply(drule crename_AndL1)
   694       apply(drule crename_AndL1)
   695 apply(simp add: fresh_prod abs_fresh fresh_atm)
   695        apply(simp add: fresh_prod abs_fresh fresh_atm)
   696 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   696       apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   697 apply(rule_tac x="Cut <a1>.M'a (x).N'a" in exI)
   697       apply(rule_tac x="Cut <a1>.M'a (x).N'a" in exI)
   698 apply(simp add: fresh_atm)[1]
   698       apply(simp add: fresh_atm)[1]
   699 apply(rule l_redu.intros)
   699       apply(rule l_redu.intros)
   700 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   700            apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   701 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   701           apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   702 apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   702          apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   703 apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   703         apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   704 apply(simp add: fresh_atm)
   704        apply(simp add: fresh_atm)
   705 apply(simp add: fresh_atm)
   705       apply(simp add: fresh_atm)
   706 (* LAnd2 *)
   706     (* LAnd2 *)
   707 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   707      apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   708 apply(drule sym)
   708      apply(drule sym)
   709 apply(drule crename_Cut)
   709      apply(drule crename_Cut)
   710 apply(simp add: fresh_prod abs_fresh fresh_atm)
   710        apply(simp add: fresh_prod abs_fresh fresh_atm)
   711 apply(simp add: fresh_prod abs_fresh fresh_atm)
   711       apply(simp add: fresh_prod abs_fresh fresh_atm)
   712 apply(auto)[1]
   712      apply(auto)[1]
   713 apply(drule crename_AndR)
   713      apply(drule crename_AndR)
   714 apply(simp add: fresh_prod abs_fresh fresh_atm)
   714         apply(simp add: fresh_prod abs_fresh fresh_atm)
   715 apply(simp add: fresh_prod abs_fresh fresh_atm)
   715        apply(simp add: fresh_prod abs_fresh fresh_atm)
   716 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   716       apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   717 apply(drule crename_AndL2)
   717      apply(drule crename_AndL2)
   718 apply(simp add: fresh_prod abs_fresh fresh_atm)
   718       apply(simp add: fresh_prod abs_fresh fresh_atm)
   719 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   719      apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   720 apply(rule_tac x="Cut <a2>.N'b (x).N'a" in exI)
   720      apply(rule_tac x="Cut <a2>.N'b (x).N'a" in exI)
   721 apply(simp add: fresh_atm)[1]
   721      apply(simp add: fresh_atm)[1]
   722 apply(rule l_redu.intros)
   722      apply(rule l_redu.intros)
   723 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   723           apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   724 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   724          apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   725 apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   725         apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   726 apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   726        apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   727 apply(simp add: fresh_atm)
   727       apply(simp add: fresh_atm)
   728 apply(simp add: fresh_atm)
   728      apply(simp add: fresh_atm)
   729 (* LOr1 *)
   729     (* LOr1 *)
   730 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   730     apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   731 apply(drule sym)
   731     apply(drule sym)
   732 apply(drule crename_Cut)
   732     apply(drule crename_Cut)
   733 apply(simp add: fresh_prod abs_fresh fresh_atm)
   733       apply(simp add: fresh_prod abs_fresh fresh_atm)
   734 apply(simp add: fresh_prod abs_fresh fresh_atm)
   734      apply(simp add: fresh_prod abs_fresh fresh_atm)
   735 apply(auto)[1]
   735     apply(auto)[1]
   736 apply(drule crename_OrL)
   736     apply(drule crename_OrL)
   737 apply(simp add: fresh_prod abs_fresh fresh_atm)
   737       apply(simp add: fresh_prod abs_fresh fresh_atm)
   738 apply(simp add: fresh_prod abs_fresh fresh_atm)
   738      apply(simp add: fresh_prod abs_fresh fresh_atm)
   739 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   739     apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   740 apply(drule crename_OrR1)
   740     apply(drule crename_OrR1)
   741 apply(simp add: fresh_prod abs_fresh fresh_atm)
   741       apply(simp add: fresh_prod abs_fresh fresh_atm)
   742 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   742      apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   743 apply(auto)
   743     apply(auto)
   744 apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
   744     apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
   745 apply(rule conjI)
   745     apply(rule conjI)
   746 apply(simp add: abs_fresh fresh_atm)[1]
   746      apply(simp add: abs_fresh fresh_atm)[1]
   747 apply(rule l_redu.intros)
   747     apply(rule l_redu.intros)
   748 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   748          apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   749 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   749         apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   750 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   750        apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   751 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   751       apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   752 apply(simp add: fresh_atm)
   752      apply(simp add: fresh_atm)
   753 apply(simp add: fresh_atm)
   753     apply(simp add: fresh_atm)
   754 (* LOr2 *)
   754     (* LOr2 *)
   755 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   755    apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   756 apply(drule sym)
   756    apply(drule sym)
   757 apply(drule crename_Cut)
   757    apply(drule crename_Cut)
   758 apply(simp add: fresh_prod abs_fresh fresh_atm)
   758      apply(simp add: fresh_prod abs_fresh fresh_atm)
   759 apply(simp add: fresh_prod abs_fresh fresh_atm)
   759     apply(simp add: fresh_prod abs_fresh fresh_atm)
   760 apply(auto)[1]
   760    apply(auto)[1]
   761 apply(drule crename_OrL)
   761    apply(drule crename_OrL)
   762 apply(simp add: fresh_prod abs_fresh fresh_atm)
   762      apply(simp add: fresh_prod abs_fresh fresh_atm)
   763 apply(simp add: fresh_prod abs_fresh fresh_atm)
   763     apply(simp add: fresh_prod abs_fresh fresh_atm)
   764 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   764    apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   765 apply(drule crename_OrR2)
   765    apply(drule crename_OrR2)
   766 apply(simp add: fresh_prod abs_fresh fresh_atm)
   766      apply(simp add: fresh_prod abs_fresh fresh_atm)
   767 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   767     apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   768 apply(auto)
   768    apply(auto)
   769 apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
   769    apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
   770 apply(rule conjI)
   770    apply(rule conjI)
   771 apply(simp add: abs_fresh fresh_atm)[1]
   771     apply(simp add: abs_fresh fresh_atm)[1]
   772 apply(rule l_redu.intros)
   772    apply(rule l_redu.intros)
   773 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   773         apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
   774 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   774        apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
   775 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   775       apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   776 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   776      apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
   777 apply(simp add: fresh_atm)
   777     apply(simp add: fresh_atm)
   778 apply(simp add: fresh_atm)
   778    apply(simp add: fresh_atm)
   779 (* ImpL *)
   779     (* ImpL *)
   780 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   780   apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
   781 apply(drule sym)
   781   apply(drule sym)
   782 apply(drule crename_Cut)
   782   apply(drule crename_Cut)
   783 apply(simp add: fresh_prod abs_fresh fresh_atm)
   783     apply(simp add: fresh_prod abs_fresh fresh_atm)
   784 apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
   784    apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
   785 apply(auto)[1]
   785   apply(auto)[1]
   786 apply(drule crename_ImpL)
   786   apply(drule crename_ImpL)
   787 apply(simp add: fresh_prod abs_fresh fresh_atm)
   787     apply(simp add: fresh_prod abs_fresh fresh_atm)
   788 apply(simp add: fresh_prod abs_fresh fresh_atm)
   788    apply(simp add: fresh_prod abs_fresh fresh_atm)
   789 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   789   apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   790 apply(drule crename_ImpR)
   790   apply(drule crename_ImpR)
   791 apply(simp add: fresh_prod abs_fresh fresh_atm)
   791      apply(simp add: fresh_prod abs_fresh fresh_atm)
   792 apply(simp add: fresh_prod abs_fresh fresh_atm)
   792     apply(simp add: fresh_prod abs_fresh fresh_atm)
   793 apply(simp)
   793    apply(simp)
   794 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   794   apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
   795 apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
   795   apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
   796 apply(rule conjI)
   796   apply(rule conjI)
   797 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   797    apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   798 apply(rule l_redu.intros)
   798   apply(rule l_redu.intros)
   799 apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
   799        apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
   800 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   800       apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   801 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
   801      apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
   802 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   802     apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   803 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   803    apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   804 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   804   apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
   805 done
   805   done
   806 
   806 
   807 lemma crename_aredu:
   807 lemma crename_aredu:
   808   assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>a M'" "a\<noteq>b"
   808   assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>a M'" "a\<noteq>b"
   809   shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>a M0"
   809   shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>a M0"
   810 using a
   810   using a
   811 apply(nominal_induct "M[a\<turnstile>c>b]" M' avoiding: M a b rule: a_redu.strong_induct)
   811   apply(nominal_induct "M[a\<turnstile>c>b]" M' avoiding: M a b rule: a_redu.strong_induct)
   812 apply(drule  crename_lredu)
   812                   apply(drule  crename_lredu)
   813 apply(blast)
   813                   apply(blast)
   814 apply(drule  crename_credu)
   814                  apply(drule  crename_credu)
   815 apply(blast)
   815                  apply(blast)
   816 (* Cut *)
   816     (* Cut *)
   817 apply(drule sym)
   817                 apply(drule sym)
   818 apply(drule crename_Cut)
   818                 apply(drule crename_Cut)
   819 apply(simp)
   819                   apply(simp)
   820 apply(simp)
   820                  apply(simp)
   821 apply(auto)[1]
   821                 apply(auto)[1]
   822 apply(drule_tac x="M'a" in meta_spec)
   822                 apply(drule_tac x="M'a" in meta_spec)
   823 apply(drule_tac x="aa" in meta_spec)
   823                 apply(drule_tac x="aa" in meta_spec)
   824 apply(drule_tac x="b" in meta_spec)
   824                 apply(drule_tac x="b" in meta_spec)
   825 apply(auto)[1]
   825                 apply(auto)[1]
   826 apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
   826                 apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
   827 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   827                 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   828 apply(rule conjI)
   828                 apply(rule conjI)
   829 apply(rule trans)
   829                  apply(rule trans)
   830 apply(rule crename.simps)
   830                   apply(rule crename.simps)
   831 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   831                    apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   832 apply(drule crename_fresh_interesting2)
   832                   apply(drule crename_fresh_interesting2)
   833 apply(simp add: fresh_a_redu)
   833                   apply(simp add: fresh_a_redu)
   834 apply(simp)
   834                  apply(simp)
   835 apply(auto)[1]
   835                 apply(auto)[1]
   836 apply(drule sym)
   836                apply(drule sym)
   837 apply(drule crename_Cut)
   837                apply(drule crename_Cut)
   838 apply(simp)
   838                  apply(simp)
   839 apply(simp)
   839                 apply(simp)
   840 apply(auto)[1]
   840                apply(auto)[1]
   841 apply(drule_tac x="N'a" in meta_spec)
   841                apply(drule_tac x="N'a" in meta_spec)
   842 apply(drule_tac x="aa" in meta_spec)
   842                apply(drule_tac x="aa" in meta_spec)
   843 apply(drule_tac x="b" in meta_spec)
   843                apply(drule_tac x="b" in meta_spec)
   844 apply(auto)[1]
   844                apply(auto)[1]
   845 apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
   845                apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
   846 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   846                apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   847 apply(rule conjI)
   847                apply(rule conjI)
   848 apply(rule trans)
   848                 apply(rule trans)
   849 apply(rule crename.simps)
   849                  apply(rule crename.simps)
   850 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
   850                   apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
   851 apply(drule crename_fresh_interesting1)
   851                   apply(drule crename_fresh_interesting1)
   852 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   852                    apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   853 apply(simp add: fresh_a_redu)
   853                   apply(simp add: fresh_a_redu)
   854 apply(simp)
   854                  apply(simp)
   855 apply(simp)
   855                 apply(simp)
   856 apply(auto)[1]
   856                apply(auto)[1]
   857 (* NotL *)
   857     (* NotL *)
   858 apply(drule sym)
   858               apply(drule sym)
   859 apply(drule crename_NotL)
   859               apply(drule crename_NotL)
   860 apply(simp)
   860                apply(simp)
   861 apply(auto)[1]
   861               apply(auto)[1]
   862 apply(drule_tac x="N'" in meta_spec)
   862               apply(drule_tac x="N'" in meta_spec)
   863 apply(drule_tac x="aa" in meta_spec)
   863               apply(drule_tac x="aa" in meta_spec)
   864 apply(drule_tac x="b" in meta_spec)
   864               apply(drule_tac x="b" in meta_spec)
   865 apply(auto)[1]
   865               apply(auto)[1]
   866 apply(rule_tac x="NotL <a>.M0 x" in exI)
   866               apply(rule_tac x="NotL <a>.M0 x" in exI)
   867 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   867               apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   868 apply(auto)[1]
   868               apply(auto)[1]
   869 (* NotR *)
   869     (* NotR *)
   870 apply(drule sym)
   870              apply(drule sym)
   871 apply(frule crename_NotR_aux)
   871              apply(frule crename_NotR_aux)
   872 apply(erule disjE)
   872              apply(erule disjE)
   873 apply(auto)[1]
   873               apply(auto)[1]
   874 apply(drule crename_NotR')
   874              apply(drule crename_NotR')
   875 apply(simp)
   875                apply(simp)
   876 apply(simp add: fresh_atm)
   876               apply(simp add: fresh_atm)
   877 apply(erule disjE)
   877              apply(erule disjE)
   878 apply(auto)[1]
   878               apply(auto)[1]
   879 apply(drule_tac x="N'" in meta_spec)
   879               apply(drule_tac x="N'" in meta_spec)
   880 apply(drule_tac x="aa" in meta_spec)
   880               apply(drule_tac x="aa" in meta_spec)
   881 apply(drule_tac x="b" in meta_spec)
   881               apply(drule_tac x="b" in meta_spec)
   882 apply(auto)[1]
   882               apply(auto)[1]
   883 apply(rule_tac x="NotR (x).M0 a" in exI)
   883               apply(rule_tac x="NotR (x).M0 a" in exI)
   884 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   884               apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   885 apply(auto)[1]
   885               apply(auto)[1]
   886 apply(auto)[1]
   886              apply(auto)[1]
   887 apply(drule_tac x="N'" in meta_spec)
   887              apply(drule_tac x="N'" in meta_spec)
   888 apply(drule_tac x="aa" in meta_spec)
   888              apply(drule_tac x="aa" in meta_spec)
   889 apply(drule_tac x="a" in meta_spec)
   889              apply(drule_tac x="a" in meta_spec)
   890 apply(auto)[1]
   890              apply(auto)[1]
   891 apply(rule_tac x="NotR (x).M0 aa" in exI)
   891              apply(rule_tac x="NotR (x).M0 aa" in exI)
   892 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   892              apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   893 apply(auto)[1]
   893              apply(auto)[1]
   894 (* AndR *)
   894     (* AndR *)
   895 apply(drule sym)
   895             apply(drule sym)
   896 apply(frule crename_AndR_aux)
   896             apply(frule crename_AndR_aux)
   897 apply(erule disjE)
   897             apply(erule disjE)
   898 apply(auto)[1]
   898              apply(auto)[1]
   899 apply(drule crename_AndR')
   899             apply(drule crename_AndR')
   900 apply(simp add: fresh_prod fresh_atm)
   900                apply(simp add: fresh_prod fresh_atm)
   901 apply(simp add: fresh_atm)
   901               apply(simp add: fresh_atm)
   902 apply(simp add: fresh_atm)
   902              apply(simp add: fresh_atm)
   903 apply(erule disjE)
   903             apply(erule disjE)
   904 apply(auto)[1]
   904              apply(auto)[1]
   905 apply(drule_tac x="M'a" in meta_spec)
   905              apply(drule_tac x="M'a" in meta_spec)
   906 apply(drule_tac x="aa" in meta_spec)
   906              apply(drule_tac x="aa" in meta_spec)
   907 apply(drule_tac x="ba" in meta_spec)
   907              apply(drule_tac x="ba" in meta_spec)
   908 apply(auto)[1]
   908              apply(auto)[1]
   909 apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
   909              apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
   910 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   910              apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   911 apply(auto)[1]
   911              apply(auto)[1]
   912 apply(rule trans)
   912              apply(rule trans)
   913 apply(rule crename.simps)
   913               apply(rule crename.simps)
   914 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   914                 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   915 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   915                apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   916 apply(auto intro: fresh_a_redu)[1]
   916                apply(auto intro: fresh_a_redu)[1]
   917 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   917               apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   918 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   918              apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   919 apply(auto)[1]
   919             apply(auto)[1]
   920 apply(drule_tac x="M'a" in meta_spec)
   920             apply(drule_tac x="M'a" in meta_spec)
   921 apply(drule_tac x="aa" in meta_spec)
   921             apply(drule_tac x="aa" in meta_spec)
   922 apply(drule_tac x="c" in meta_spec)
   922             apply(drule_tac x="c" in meta_spec)
   923 apply(auto)[1]
   923             apply(auto)[1]
   924 apply(rule_tac x="AndR <a>.M0 <b>.N' aa" in exI)
   924             apply(rule_tac x="AndR <a>.M0 <b>.N' aa" in exI)
   925 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   925             apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   926 apply(auto)[1]
   926             apply(auto)[1]
   927 apply(rule trans)
   927             apply(rule trans)
   928 apply(rule crename.simps)
   928              apply(rule crename.simps)
   929 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   929                apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   930 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   930               apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   931 apply(auto intro: fresh_a_redu)[1]
   931               apply(auto intro: fresh_a_redu)[1]
   932 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   932              apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   933 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   933             apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   934 apply(drule sym)
   934            apply(drule sym)
   935 apply(frule crename_AndR_aux)
   935            apply(frule crename_AndR_aux)
   936 apply(erule disjE)
   936            apply(erule disjE)
   937 apply(auto)[1]
   937             apply(auto)[1]
   938 apply(drule crename_AndR')
   938            apply(drule crename_AndR')
   939 apply(simp add: fresh_prod fresh_atm)
   939               apply(simp add: fresh_prod fresh_atm)
   940 apply(simp add: fresh_atm)
   940              apply(simp add: fresh_atm)
   941 apply(simp add: fresh_atm)
   941             apply(simp add: fresh_atm)
   942 apply(erule disjE)
   942            apply(erule disjE)
   943 apply(auto)[1]
   943             apply(auto)[1]
   944 apply(drule_tac x="N'a" in meta_spec)
   944             apply(drule_tac x="N'a" in meta_spec)
   945 apply(drule_tac x="aa" in meta_spec)
   945             apply(drule_tac x="aa" in meta_spec)
   946 apply(drule_tac x="ba" in meta_spec)
   946             apply(drule_tac x="ba" in meta_spec)
   947 apply(auto)[1]
   947             apply(auto)[1]
   948 apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
   948             apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
   949 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   949             apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   950 apply(auto)[1]
   950             apply(auto)[1]
   951 apply(rule trans)
   951             apply(rule trans)
   952 apply(rule crename.simps)
   952              apply(rule crename.simps)
   953 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
   953                apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
   954 apply(auto intro: fresh_a_redu)[1]
   954                apply(auto intro: fresh_a_redu)[1]
   955 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   955               apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   956 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   956              apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   957 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   957             apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   958 apply(auto)[1]
   958            apply(auto)[1]
   959 apply(drule_tac x="N'a" in meta_spec)
   959            apply(drule_tac x="N'a" in meta_spec)
   960 apply(drule_tac x="aa" in meta_spec)
   960            apply(drule_tac x="aa" in meta_spec)
   961 apply(drule_tac x="c" in meta_spec)
   961            apply(drule_tac x="c" in meta_spec)
   962 apply(auto)[1]
   962            apply(auto)[1]
   963 apply(rule_tac x="AndR <a>.M' <b>.M0 aa" in exI)
   963            apply(rule_tac x="AndR <a>.M' <b>.M0 aa" in exI)
   964 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   964            apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   965 apply(auto)[1]
   965            apply(auto)[1]
   966 apply(rule trans)
   966            apply(rule trans)
   967 apply(rule crename.simps)
   967             apply(rule crename.simps)
   968 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
   968               apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
   969 apply(auto intro: fresh_a_redu)[1]
   969               apply(auto intro: fresh_a_redu)[1]
   970 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   970              apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   971 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   971             apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
   972 apply(simp)
   972            apply(simp)
   973 (* AndL1 *)
   973     (* AndL1 *)
   974 apply(drule sym)
   974           apply(drule sym)
   975 apply(drule crename_AndL1)
   975           apply(drule crename_AndL1)
   976 apply(simp)
   976            apply(simp)
   977 apply(auto)[1]
   977           apply(auto)[1]
   978 apply(drule_tac x="N'" in meta_spec)
   978           apply(drule_tac x="N'" in meta_spec)
   979 apply(drule_tac x="a" in meta_spec)
   979           apply(drule_tac x="a" in meta_spec)
   980 apply(drule_tac x="b" in meta_spec)
   980           apply(drule_tac x="b" in meta_spec)
   981 apply(auto)[1]
   981           apply(auto)[1]
   982 apply(rule_tac x="AndL1 (x).M0 y" in exI)
   982           apply(rule_tac x="AndL1 (x).M0 y" in exI)
   983 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   983           apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   984 apply(auto)[1]
   984           apply(auto)[1]
   985 (* AndL2 *)
   985     (* AndL2 *)
   986 apply(drule sym)
   986          apply(drule sym)
   987 apply(drule crename_AndL2)
   987          apply(drule crename_AndL2)
   988 apply(simp)
   988           apply(simp)
   989 apply(auto)[1]
   989          apply(auto)[1]
   990 apply(drule_tac x="N'" in meta_spec)
   990          apply(drule_tac x="N'" in meta_spec)
   991 apply(drule_tac x="a" in meta_spec)
   991          apply(drule_tac x="a" in meta_spec)
   992 apply(drule_tac x="b" in meta_spec)
   992          apply(drule_tac x="b" in meta_spec)
   993 apply(auto)[1]
   993          apply(auto)[1]
   994 apply(rule_tac x="AndL2 (x).M0 y" in exI)
   994          apply(rule_tac x="AndL2 (x).M0 y" in exI)
   995 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   995          apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
   996 apply(auto)[1]
   996          apply(auto)[1]
   997 (* OrL *)
   997     (* OrL *)
   998 apply(drule sym)
   998         apply(drule sym)
   999 apply(drule crename_OrL)
   999         apply(drule crename_OrL)
  1000 apply(simp)
  1000           apply(simp)
  1001 apply(auto simp add: fresh_atm fresh_prod)[1]
  1001           apply(auto simp add: fresh_atm fresh_prod)[1]
  1002 apply(auto simp add: fresh_atm fresh_prod)[1]
  1002          apply(auto simp add: fresh_atm fresh_prod)[1]
  1003 apply(auto)[1]
  1003         apply(auto)[1]
  1004 apply(drule_tac x="M'a" in meta_spec)
  1004         apply(drule_tac x="M'a" in meta_spec)
  1005 apply(drule_tac x="a" in meta_spec)
  1005         apply(drule_tac x="a" in meta_spec)
  1006 apply(drule_tac x="b" in meta_spec)
  1006         apply(drule_tac x="b" in meta_spec)
  1007 apply(auto)[1]
  1007         apply(auto)[1]
  1008 apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
  1008         apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
  1009 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1009         apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1010 apply(auto)[1]
  1010         apply(auto)[1]
  1011 apply(rule trans)
  1011         apply(rule trans)
  1012 apply(rule crename.simps)
  1012          apply(rule crename.simps)
  1013 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1013            apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1014 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1014           apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1015 apply(auto intro: fresh_a_redu)[1]
  1015           apply(auto intro: fresh_a_redu)[1]
  1016 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1016          apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1017 apply(simp)
  1017         apply(simp)
  1018 apply(drule sym)
  1018        apply(drule sym)
  1019 apply(drule crename_OrL)
  1019        apply(drule crename_OrL)
  1020 apply(simp)
  1020          apply(simp)
  1021 apply(auto simp add: fresh_atm fresh_prod)[1]
  1021          apply(auto simp add: fresh_atm fresh_prod)[1]
  1022 apply(auto simp add: fresh_atm fresh_prod)[1]
  1022         apply(auto simp add: fresh_atm fresh_prod)[1]
  1023 apply(auto)[1]
  1023        apply(auto)[1]
  1024 apply(drule_tac x="N'a" in meta_spec)
  1024        apply(drule_tac x="N'a" in meta_spec)
  1025 apply(drule_tac x="a" in meta_spec)
  1025        apply(drule_tac x="a" in meta_spec)
  1026 apply(drule_tac x="b" in meta_spec)
  1026        apply(drule_tac x="b" in meta_spec)
  1027 apply(auto)[1]
  1027        apply(auto)[1]
  1028 apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
  1028        apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
  1029 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1029        apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1030 apply(auto)[1]
  1030        apply(auto)[1]
  1031 apply(rule trans)
  1031        apply(rule trans)
  1032 apply(rule crename.simps)
  1032         apply(rule crename.simps)
  1033 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1033           apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1034 apply(auto intro: fresh_a_redu)[1]
  1034           apply(auto intro: fresh_a_redu)[1]
  1035 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1035          apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1036 apply(simp)
  1036         apply(simp)
  1037 apply(simp)
  1037        apply(simp)
  1038 (* OrR1 *)
  1038     (* OrR1 *)
  1039 apply(drule sym)
  1039       apply(drule sym)
  1040 apply(frule crename_OrR1_aux)
  1040       apply(frule crename_OrR1_aux)
  1041 apply(erule disjE)
  1041       apply(erule disjE)
  1042 apply(auto)[1]
  1042        apply(auto)[1]
  1043 apply(drule crename_OrR1')
  1043       apply(drule crename_OrR1')
  1044 apply(simp)
  1044         apply(simp)
  1045 apply(simp add: fresh_atm)
  1045        apply(simp add: fresh_atm)
  1046 apply(erule disjE)
  1046       apply(erule disjE)
  1047 apply(auto)[1]
  1047        apply(auto)[1]
  1048 apply(drule_tac x="N'" in meta_spec)
  1048        apply(drule_tac x="N'" in meta_spec)
  1049 apply(drule_tac x="aa" in meta_spec)
  1049        apply(drule_tac x="aa" in meta_spec)
  1050 apply(drule_tac x="ba" in meta_spec)
  1050        apply(drule_tac x="ba" in meta_spec)
  1051 apply(auto)[1]
  1051        apply(auto)[1]
  1052 apply(rule_tac x="OrR1 <a>.M0 b" in exI)
  1052        apply(rule_tac x="OrR1 <a>.M0 b" in exI)
  1053 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1053        apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1054 apply(auto)[1]
  1054        apply(auto)[1]
  1055 apply(auto)[1]
  1055       apply(auto)[1]
  1056 apply(drule_tac x="N'" in meta_spec)
  1056       apply(drule_tac x="N'" in meta_spec)
  1057 apply(drule_tac x="aa" in meta_spec)
  1057       apply(drule_tac x="aa" in meta_spec)
  1058 apply(drule_tac x="b" in meta_spec)
  1058       apply(drule_tac x="b" in meta_spec)
  1059 apply(auto)[1]
  1059       apply(auto)[1]
  1060 apply(rule_tac x="OrR1 <a>.M0 aa" in exI)
  1060       apply(rule_tac x="OrR1 <a>.M0 aa" in exI)
  1061 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1061       apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1062 apply(auto)[1]
  1062       apply(auto)[1]
  1063 (* OrR2 *)
  1063     (* OrR2 *)
  1064 apply(drule sym)
  1064      apply(drule sym)
  1065 apply(frule crename_OrR2_aux)
  1065      apply(frule crename_OrR2_aux)
  1066 apply(erule disjE)
  1066      apply(erule disjE)
  1067 apply(auto)[1]
  1067       apply(auto)[1]
  1068 apply(drule crename_OrR2')
  1068      apply(drule crename_OrR2')
  1069 apply(simp)
  1069        apply(simp)
  1070 apply(simp add: fresh_atm)
  1070       apply(simp add: fresh_atm)
  1071 apply(erule disjE)
  1071      apply(erule disjE)
  1072 apply(auto)[1]
  1072       apply(auto)[1]
  1073 apply(drule_tac x="N'" in meta_spec)
  1073       apply(drule_tac x="N'" in meta_spec)
  1074 apply(drule_tac x="aa" in meta_spec)
  1074       apply(drule_tac x="aa" in meta_spec)
  1075 apply(drule_tac x="ba" in meta_spec)
  1075       apply(drule_tac x="ba" in meta_spec)
  1076 apply(auto)[1]
  1076       apply(auto)[1]
  1077 apply(rule_tac x="OrR2 <a>.M0 b" in exI)
  1077       apply(rule_tac x="OrR2 <a>.M0 b" in exI)
  1078 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1078       apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1079 apply(auto)[1]
  1079       apply(auto)[1]
  1080 apply(auto)[1]
  1080      apply(auto)[1]
  1081 apply(drule_tac x="N'" in meta_spec)
  1081      apply(drule_tac x="N'" in meta_spec)
  1082 apply(drule_tac x="aa" in meta_spec)
  1082      apply(drule_tac x="aa" in meta_spec)
  1083 apply(drule_tac x="b" in meta_spec)
  1083      apply(drule_tac x="b" in meta_spec)
  1084 apply(auto)[1]
  1084      apply(auto)[1]
  1085 apply(rule_tac x="OrR2 <a>.M0 aa" in exI)
  1085      apply(rule_tac x="OrR2 <a>.M0 aa" in exI)
  1086 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1086      apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1087 apply(auto)[1]
  1087      apply(auto)[1]
  1088 (* ImpL *)
  1088     (* ImpL *)
  1089 apply(drule sym)
  1089     apply(drule sym)
  1090 apply(drule crename_ImpL)
  1090     apply(drule crename_ImpL)
  1091 apply(simp)
  1091       apply(simp)
  1092 apply(simp)
  1092      apply(simp)
  1093 apply(auto)[1]
  1093     apply(auto)[1]
  1094 apply(drule_tac x="M'a" in meta_spec)
  1094     apply(drule_tac x="M'a" in meta_spec)
  1095 apply(drule_tac x="aa" in meta_spec)
  1095     apply(drule_tac x="aa" in meta_spec)
  1096 apply(drule_tac x="b" in meta_spec)
  1096     apply(drule_tac x="b" in meta_spec)
  1097 apply(auto)[1]
  1097     apply(auto)[1]
  1098 apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
  1098     apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
  1099 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1099     apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1100 apply(auto)[1]
  1100     apply(auto)[1]
  1101 apply(rule trans)
  1101     apply(rule trans)
  1102 apply(rule crename.simps)
  1102      apply(rule crename.simps)
  1103 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1103       apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1104 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1104      apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1105 apply(auto intro: fresh_a_redu)[1]
  1105      apply(auto intro: fresh_a_redu)[1]
  1106 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1106     apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1107 apply(drule sym)
  1107    apply(drule sym)
  1108 apply(drule crename_ImpL)
  1108    apply(drule crename_ImpL)
  1109 apply(simp)
  1109      apply(simp)
  1110 apply(simp)
  1110     apply(simp)
  1111 apply(auto)[1]
  1111    apply(auto)[1]
  1112 apply(drule_tac x="N'a" in meta_spec)
  1112    apply(drule_tac x="N'a" in meta_spec)
  1113 apply(drule_tac x="aa" in meta_spec)
  1113    apply(drule_tac x="aa" in meta_spec)
  1114 apply(drule_tac x="b" in meta_spec)
  1114    apply(drule_tac x="b" in meta_spec)
  1115 apply(auto)[1]
  1115    apply(auto)[1]
  1116 apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
  1116    apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
  1117 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1117    apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1118 apply(auto)[1]
  1118    apply(auto)[1]
  1119 apply(rule trans)
  1119    apply(rule trans)
  1120 apply(rule crename.simps)
  1120     apply(rule crename.simps)
  1121 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1121      apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1122 apply(auto intro: fresh_a_redu)[1]
  1122      apply(auto intro: fresh_a_redu)[1]
  1123 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1123     apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1124 apply(simp)
  1124    apply(simp)
  1125 (* ImpR *)
  1125     (* ImpR *)
  1126 apply(drule sym)
  1126   apply(drule sym)
  1127 apply(frule crename_ImpR_aux)
  1127   apply(frule crename_ImpR_aux)
  1128 apply(erule disjE)
  1128   apply(erule disjE)
  1129 apply(auto)[1]
  1129    apply(auto)[1]
  1130 apply(drule crename_ImpR')
  1130   apply(drule crename_ImpR')
  1131 apply(simp)
  1131      apply(simp)
  1132 apply(simp add: fresh_atm)
  1132     apply(simp add: fresh_atm)
  1133 apply(simp add: fresh_atm)
  1133    apply(simp add: fresh_atm)
  1134 apply(erule disjE)
  1134   apply(erule disjE)
  1135 apply(auto)[1]
  1135    apply(auto)[1]
  1136 apply(drule_tac x="N'" in meta_spec)
  1136    apply(drule_tac x="N'" in meta_spec)
  1137 apply(drule_tac x="aa" in meta_spec)
  1137    apply(drule_tac x="aa" in meta_spec)
  1138 apply(drule_tac x="ba" in meta_spec)
  1138    apply(drule_tac x="ba" in meta_spec)
  1139 apply(auto)[1]
  1139    apply(auto)[1]
  1140 apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
  1140    apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
  1141 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1141    apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1142 apply(auto)[1]
  1142    apply(auto)[1]
  1143 apply(auto)[1]
  1143   apply(auto)[1]
  1144 apply(drule_tac x="N'" in meta_spec)
  1144   apply(drule_tac x="N'" in meta_spec)
  1145 apply(drule_tac x="aa" in meta_spec)
  1145   apply(drule_tac x="aa" in meta_spec)
  1146 apply(drule_tac x="b" in meta_spec)
  1146   apply(drule_tac x="b" in meta_spec)
  1147 apply(auto)[1]
  1147   apply(auto)[1]
  1148 apply(rule_tac x="ImpR (x).<a>.M0 aa" in exI)
  1148   apply(rule_tac x="ImpR (x).<a>.M0 aa" in exI)
  1149 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1149   apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1150 apply(auto)[1]
  1150   apply(auto)[1]
  1151 done
  1151   done
  1152 
  1152 
  1153 lemma SNa_preserved_renaming1:
  1153 lemma SNa_preserved_renaming1:
  1154   assumes a: "SNa M"
  1154   assumes a: "SNa M"
  1155   shows "SNa (M[a\<turnstile>c>b])"
  1155   shows "SNa (M[a\<turnstile>c>b])"
  1156 using a
  1156   using a
  1157 apply(induct rule: SNa_induct)
  1157   apply(induct rule: SNa_induct)
  1158 apply(case_tac "a=b")
  1158   apply(case_tac "a=b")
  1159 apply(simp add: crename_id)
  1159    apply(simp add: crename_id)
  1160 apply(rule SNaI)
  1160   apply(rule SNaI)
  1161 apply(drule crename_aredu)
  1161   apply(drule crename_aredu)
  1162 apply(blast)+
  1162    apply(blast)+
  1163 done
  1163   done
  1164 
  1164 
  1165 lemma nrename_interesting1:
  1165 lemma nrename_interesting1:
  1166   assumes a: "distinct [x,y,z]"
  1166   assumes a: "distinct [x,y,z]"
  1167   shows "M[x\<turnstile>n>z][z\<turnstile>n>y] = M[z\<turnstile>n>y][x\<turnstile>n>y]"
  1167   shows "M[x\<turnstile>n>z][z\<turnstile>n>y] = M[z\<turnstile>n>y][x\<turnstile>n>y]"
  1168 using a
  1168   using a
  1169 apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1169   apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1170 apply(auto simp add: rename_fresh simp add: trm.inject alpha)
  1170              apply(auto simp add: rename_fresh simp add: trm.inject alpha)
  1171 apply(blast)
  1171      apply(blast)
  1172 apply(blast)
  1172     apply(blast)
  1173 apply(rotate_tac 12)
  1173    apply(rotate_tac 12)
  1174 apply(drule_tac x="x" in meta_spec)
  1174    apply(drule_tac x="x" in meta_spec)
  1175 apply(rotate_tac 15)
  1175    apply(rotate_tac 15)
  1176 apply(drule_tac x="y" in meta_spec)
  1176    apply(drule_tac x="y" in meta_spec)
  1177 apply(rotate_tac 15)
  1177    apply(rotate_tac 15)
  1178 apply(drule_tac x="z" in meta_spec)
  1178    apply(drule_tac x="z" in meta_spec)
  1179 apply(blast)
  1179    apply(blast)
  1180 apply(rotate_tac 11)
  1180   apply(rotate_tac 11)
  1181 apply(drule_tac x="x" in meta_spec)
  1181   apply(drule_tac x="x" in meta_spec)
  1182 apply(rotate_tac 14)
  1182   apply(rotate_tac 14)
  1183 apply(drule_tac x="y" in meta_spec)
  1183   apply(drule_tac x="y" in meta_spec)
  1184 apply(rotate_tac 14)
  1184   apply(rotate_tac 14)
  1185 apply(drule_tac x="z" in meta_spec)
  1185   apply(drule_tac x="z" in meta_spec)
  1186 apply(blast)
  1186   apply(blast)
  1187 done
  1187   done
  1188 
  1188 
  1189 lemma nrename_interesting2:
  1189 lemma nrename_interesting2:
  1190   assumes a: "x\<noteq>z" "x\<noteq>u" "x\<noteq>y" "z\<noteq>u" "y\<noteq>z"
  1190   assumes a: "x\<noteq>z" "x\<noteq>u" "x\<noteq>y" "z\<noteq>u" "y\<noteq>z"
  1191   shows "M[x\<turnstile>n>y][z\<turnstile>n>u] = M[z\<turnstile>n>u][x\<turnstile>n>y]"
  1191   shows "M[x\<turnstile>n>y][z\<turnstile>n>u] = M[z\<turnstile>n>u][x\<turnstile>n>y]"
  1192 using a
  1192   using a
  1193 apply(nominal_induct M avoiding: x y z u rule: trm.strong_induct)
  1193   apply(nominal_induct M avoiding: x y z u rule: trm.strong_induct)
  1194 apply(auto simp add: rename_fresh simp add: trm.inject alpha)
  1194              apply(auto simp add: rename_fresh simp add: trm.inject alpha)
  1195 done
  1195   done
  1196 
  1196 
  1197 lemma not_fic_nrename_aux:
  1197 lemma not_fic_nrename_aux:
  1198   assumes a: "fic M c" 
  1198   assumes a: "fic M c" 
  1199   shows "fic (M[x\<turnstile>n>y]) c" 
  1199   shows "fic (M[x\<turnstile>n>y]) c" 
  1200 using a
  1200   using a
  1201 apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
  1201   apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
  1202 apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
  1202              apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
  1203 done
  1203   done
  1204 
  1204 
  1205 lemma not_fic_nrename:
  1205 lemma not_fic_nrename:
  1206   assumes a: "\<not>(fic (M[x\<turnstile>n>y]) c)" 
  1206   assumes a: "\<not>(fic (M[x\<turnstile>n>y]) c)" 
  1207   shows "\<not>(fic M c)" 
  1207   shows "\<not>(fic M c)" 
  1208 using a
  1208   using a
  1209 apply(auto dest:  not_fic_nrename_aux)
  1209   apply(auto dest:  not_fic_nrename_aux)
  1210 done
  1210   done
  1211 
  1211 
  1212 lemma fin_nrename:
  1212 lemma fin_nrename:
  1213   assumes a: "fin M z" "z\<sharp>(x,y)"
  1213   assumes a: "fin M z" "z\<sharp>(x,y)"
  1214   shows "fin (M[x\<turnstile>n>y]) z" 
  1214   shows "fin (M[x\<turnstile>n>y]) z" 
  1215 using a
  1215   using a
  1216 apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1216   apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1217 apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
  1217              apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
  1218            split: if_splits)
  1218       split: if_splits)
  1219 done
  1219   done
  1220 
  1220 
  1221 lemma nrename_fresh_interesting1:
  1221 lemma nrename_fresh_interesting1:
  1222   fixes z::"name"
  1222   fixes z::"name"
  1223   assumes a: "z\<sharp>(M[x\<turnstile>n>y])" "z\<sharp>(x,y)"
  1223   assumes a: "z\<sharp>(M[x\<turnstile>n>y])" "z\<sharp>(x,y)"
  1224   shows "z\<sharp>M"
  1224   shows "z\<sharp>M"
  1225 using a
  1225   using a
  1226 apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1226   apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1227 apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp)
  1227              apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp)
  1228 done
  1228   done
  1229 
  1229 
  1230 lemma nrename_fresh_interesting2:
  1230 lemma nrename_fresh_interesting2:
  1231   fixes c::"coname"
  1231   fixes c::"coname"
  1232   assumes a: "c\<sharp>(M[x\<turnstile>n>y])" 
  1232   assumes a: "c\<sharp>(M[x\<turnstile>n>y])" 
  1233   shows "c\<sharp>M"
  1233   shows "c\<sharp>M"
  1234 using a
  1234   using a
  1235 apply(nominal_induct M avoiding: x y c rule: trm.strong_induct)
  1235   apply(nominal_induct M avoiding: x y c rule: trm.strong_induct)
  1236 apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
  1236              apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
  1237 done
  1237   done
  1238 
  1238 
  1239 lemma fin_nrename2:
  1239 lemma fin_nrename2:
  1240   assumes a: "fin (M[x\<turnstile>n>y]) z" "z\<sharp>(x,y)"
  1240   assumes a: "fin (M[x\<turnstile>n>y]) z" "z\<sharp>(x,y)"
  1241   shows "fin M z" 
  1241   shows "fin M z" 
  1242 using a
  1242   using a
  1243 apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1243   apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
  1244 apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
  1244              apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
  1245            split: if_splits)
  1245       split: if_splits)
  1246 apply(auto dest: nrename_fresh_interesting1 simp add: fresh_atm fresh_prod)
  1246         apply(auto dest: nrename_fresh_interesting1 simp add: fresh_atm fresh_prod)
  1247 done
  1247   done
  1248 
  1248 
  1249 lemma nrename_Cut:
  1249 lemma nrename_Cut:
  1250   assumes a: "R[x\<turnstile>n>y] = Cut <c>.M (z).N" "c\<sharp>(N,R)" "z\<sharp>(x,y,M,R)"
  1250   assumes a: "R[x\<turnstile>n>y] = Cut <c>.M (z).N" "c\<sharp>(N,R)" "z\<sharp>(x,y,M,R)"
  1251   shows "\<exists>M' N'. R = Cut <c>.M' (z).N' \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> z\<sharp>M'"
  1251   shows "\<exists>M' N'. R = Cut <c>.M' (z).N' \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> z\<sharp>M'"
  1252 using a
  1252   using a
  1253 apply(nominal_induct R avoiding: c y x z M N rule: trm.strong_induct)
  1253   apply(nominal_induct R avoiding: c y x z M N rule: trm.strong_induct)
  1254 apply(auto split: if_splits)
  1254              apply(auto split: if_splits)
  1255 apply(simp add: trm.inject)
  1255   apply(simp add: trm.inject)
  1256 apply(auto simp add: alpha fresh_atm)
  1256   apply(auto simp add: alpha fresh_atm)
  1257 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1257   apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1258 apply(perm_simp)
  1258   apply(perm_simp)
  1259 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1259   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1260 apply(rule_tac x="[(name,z)]\<bullet>trm2" in exI)
  1260   apply(rule_tac x="[(name,z)]\<bullet>trm2" in exI)
  1261 apply(perm_simp)
  1261   apply(perm_simp)
  1262 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1262   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1263 apply(rule conjI)
  1263   apply(rule conjI)
  1264 apply(drule sym)
  1264    apply(drule sym)
  1265 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1265    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1266 apply(simp add: eqvts calc_atm)
  1266    apply(simp add: eqvts calc_atm)
  1267 apply(auto simp add: fresh_atm)[1]
  1267   apply(auto simp add: fresh_atm)[1]
  1268 apply(drule sym)
  1268   apply(drule sym)
  1269 apply(drule sym)
  1269   apply(drule sym)
  1270 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1270   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1271 apply(simp add: eqvts calc_atm)
  1271   apply(simp add: eqvts calc_atm)
  1272 done
  1272   done
  1273 
  1273 
  1274 lemma nrename_NotR:
  1274 lemma nrename_NotR:
  1275   assumes a: "R[x\<turnstile>n>y] = NotR (z).N c" "z\<sharp>(R,x,y)" 
  1275   assumes a: "R[x\<turnstile>n>y] = NotR (z).N c" "z\<sharp>(R,x,y)" 
  1276   shows "\<exists>N'. (R = NotR (z).N' c) \<and> N'[x\<turnstile>n>y] = N" 
  1276   shows "\<exists>N'. (R = NotR (z).N' c) \<and> N'[x\<turnstile>n>y] = N" 
  1277 using a
  1277   using a
  1278 apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
  1278   apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
  1279 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1279              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1280 apply(rule_tac x="[(name,z)]\<bullet>trm" in exI)
  1280   apply(rule_tac x="[(name,z)]\<bullet>trm" in exI)
  1281 apply(perm_simp)
  1281   apply(perm_simp)
  1282 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1282   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1283 apply(drule sym)
  1283   apply(drule sym)
  1284 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1284   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1285 apply(simp add: eqvts calc_atm)
  1285   apply(simp add: eqvts calc_atm)
  1286 done
  1286   done
  1287 
  1287 
  1288 lemma nrename_NotL:
  1288 lemma nrename_NotL:
  1289   assumes a: "R[x\<turnstile>n>y] = NotL <c>.N z" "c\<sharp>R" "z\<sharp>(x,y)"
  1289   assumes a: "R[x\<turnstile>n>y] = NotL <c>.N z" "c\<sharp>R" "z\<sharp>(x,y)"
  1290   shows "\<exists>N'. (R = NotL <c>.N' z) \<and> N'[x\<turnstile>n>y] = N" 
  1290   shows "\<exists>N'. (R = NotL <c>.N' z) \<and> N'[x\<turnstile>n>y] = N" 
  1291 using a
  1291   using a
  1292 apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
  1292   apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
  1293 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1293              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1294 apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
  1294   apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
  1295 apply(perm_simp)
  1295   apply(perm_simp)
  1296 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1296   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1297 apply(drule sym)
  1297   apply(drule sym)
  1298 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1298   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1299 apply(simp add: eqvts calc_atm)
  1299   apply(simp add: eqvts calc_atm)
  1300 done
  1300   done
  1301 
  1301 
  1302 lemma nrename_NotL':
  1302 lemma nrename_NotL':
  1303   assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" "c\<sharp>R" "x\<noteq>y" 
  1303   assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" "c\<sharp>R" "x\<noteq>y" 
  1304   shows "(\<exists>N'. (R = NotL <c>.N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = NotL <c>.N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
  1304   shows "(\<exists>N'. (R = NotL <c>.N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = NotL <c>.N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
  1305 using a
  1305   using a
  1306 apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
  1306   apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
  1307 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1307              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1308 apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
  1308    apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
  1309 apply(perm_simp)
  1309    apply(perm_simp)
  1310 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1310    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1311 apply(drule sym)
  1311    apply(drule sym)
  1312 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1312    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1313 apply(simp add: eqvts calc_atm)
  1313    apply(simp add: eqvts calc_atm)
  1314 apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
  1314   apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
  1315 apply(perm_simp)
  1315   apply(perm_simp)
  1316 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1316   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1317 apply(drule sym)
  1317   apply(drule sym)
  1318 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1318   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1319 apply(simp add: eqvts calc_atm)
  1319   apply(simp add: eqvts calc_atm)
  1320 done
  1320   done
  1321 
  1321 
  1322 lemma nrename_NotL_aux:
  1322 lemma nrename_NotL_aux:
  1323   assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" 
  1323   assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" 
  1324   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1324   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1325 using a
  1325   using a
  1326 apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
  1326   apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
  1327 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1327              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1328 done
  1328   done
  1329 
  1329 
  1330 lemma nrename_AndL1:
  1330 lemma nrename_AndL1:
  1331   assumes a: "R[x\<turnstile>n>y] = AndL1 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
  1331   assumes a: "R[x\<turnstile>n>y] = AndL1 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
  1332   shows "\<exists>N'. (R = AndL1 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
  1332   shows "\<exists>N'. (R = AndL1 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
  1333 using a
  1333   using a
  1334 apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
  1334   apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
  1335 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1335              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1336 apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
  1336   apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
  1337 apply(perm_simp)
  1337   apply(perm_simp)
  1338 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1338   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1339 apply(drule sym)
  1339   apply(drule sym)
  1340 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1340   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1341 apply(simp add: eqvts calc_atm)
  1341   apply(simp add: eqvts calc_atm)
  1342 done
  1342   done
  1343 
  1343 
  1344 lemma nrename_AndL1':
  1344 lemma nrename_AndL1':
  1345   assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
  1345   assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
  1346   shows "(\<exists>N'. (R = AndL1 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL1 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
  1346   shows "(\<exists>N'. (R = AndL1 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL1 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
  1347 using a
  1347   using a
  1348 apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1348   apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1349 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1349              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1350 apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1350    apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1351 apply(perm_simp)
  1351    apply(perm_simp)
  1352 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1352    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1353 apply(drule sym)
  1353    apply(drule sym)
  1354 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1354    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1355 apply(simp add: eqvts calc_atm)
  1355    apply(simp add: eqvts calc_atm)
  1356 apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1356   apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1357 apply(perm_simp)
  1357   apply(perm_simp)
  1358 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1358   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1359 apply(drule sym)
  1359   apply(drule sym)
  1360 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1360   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1361 apply(simp add: eqvts calc_atm)
  1361   apply(simp add: eqvts calc_atm)
  1362 done
  1362   done
  1363 
  1363 
  1364 lemma nrename_AndL1_aux:
  1364 lemma nrename_AndL1_aux:
  1365   assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" 
  1365   assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" 
  1366   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1366   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1367 using a
  1367   using a
  1368 apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1368   apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1369 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1369              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1370 done
  1370   done
  1371 
  1371 
  1372 lemma nrename_AndL2:
  1372 lemma nrename_AndL2:
  1373   assumes a: "R[x\<turnstile>n>y] = AndL2 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
  1373   assumes a: "R[x\<turnstile>n>y] = AndL2 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
  1374   shows "\<exists>N'. (R = AndL2 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
  1374   shows "\<exists>N'. (R = AndL2 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
  1375 using a
  1375   using a
  1376 apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
  1376   apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
  1377 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1377              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1378 apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
  1378   apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
  1379 apply(perm_simp)
  1379   apply(perm_simp)
  1380 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1380   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1381 apply(drule sym)
  1381   apply(drule sym)
  1382 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1382   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1383 apply(simp add: eqvts calc_atm)
  1383   apply(simp add: eqvts calc_atm)
  1384 done
  1384   done
  1385 
  1385 
  1386 lemma nrename_AndL2':
  1386 lemma nrename_AndL2':
  1387   assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
  1387   assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
  1388   shows "(\<exists>N'. (R = AndL2 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL2 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
  1388   shows "(\<exists>N'. (R = AndL2 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL2 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
  1389 using a
  1389   using a
  1390 apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1390   apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1391 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1391              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1392 apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1392    apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1393 apply(perm_simp)
  1393    apply(perm_simp)
  1394 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1394    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1395 apply(drule sym)
  1395    apply(drule sym)
  1396 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1396    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1397 apply(simp add: eqvts calc_atm)
  1397    apply(simp add: eqvts calc_atm)
  1398 apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1398   apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
  1399 apply(perm_simp)
  1399   apply(perm_simp)
  1400 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1400   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1401 apply(drule sym)
  1401   apply(drule sym)
  1402 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1402   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1403 apply(simp add: eqvts calc_atm)
  1403   apply(simp add: eqvts calc_atm)
  1404 done
  1404   done
  1405 
  1405 
  1406 lemma nrename_AndL2_aux:
  1406 lemma nrename_AndL2_aux:
  1407   assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" 
  1407   assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" 
  1408   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1408   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1409 using a
  1409   using a
  1410 apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1410   apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
  1411 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1411              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1412 done
  1412   done
  1413 
  1413 
  1414 lemma nrename_AndR:
  1414 lemma nrename_AndR:
  1415   assumes a: "R[x\<turnstile>n>y] = AndR <c>.M <d>.N e" "c\<sharp>(d,e,N,R)" "d\<sharp>(c,e,M,R)" 
  1415   assumes a: "R[x\<turnstile>n>y] = AndR <c>.M <d>.N e" "c\<sharp>(d,e,N,R)" "d\<sharp>(c,e,M,R)" 
  1416   shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
  1416   shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
  1417 using a
  1417   using a
  1418 apply(nominal_induct R avoiding: x y c d e M N rule: trm.strong_induct)
  1418   apply(nominal_induct R avoiding: x y c d e M N rule: trm.strong_induct)
  1419 apply(auto split: if_splits simp add: trm.inject alpha)
  1419              apply(auto split: if_splits simp add: trm.inject alpha)
  1420 apply(simp add: fresh_atm fresh_prod)
  1420     apply(simp add: fresh_atm fresh_prod)
  1421 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
  1421    apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
  1422 apply(perm_simp)
  1422    apply(perm_simp)
  1423 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1423    apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1424 apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
  1424   apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
  1425 apply(perm_simp)
  1425   apply(perm_simp)
  1426 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1426   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1427 apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
  1427   apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
  1428 apply(perm_simp)
  1428   apply(perm_simp)
  1429 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1429   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1430 apply(drule sym)
  1430    apply(drule sym)
  1431 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1431    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1432 apply(simp add: eqvts calc_atm)
  1432    apply(simp add: eqvts calc_atm)
  1433 apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
  1433   apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
  1434 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1434   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1435 apply(simp add: eqvts calc_atm)
  1435   apply(simp add: eqvts calc_atm)
  1436 done
  1436   done
  1437 
  1437 
  1438 lemma nrename_OrR1:
  1438 lemma nrename_OrR1:
  1439   assumes a: "R[x\<turnstile>n>y] = OrR1 <c>.N d" "c\<sharp>(R,d)" 
  1439   assumes a: "R[x\<turnstile>n>y] = OrR1 <c>.N d" "c\<sharp>(R,d)" 
  1440   shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
  1440   shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
  1441 using a
  1441   using a
  1442 apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
  1442   apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
  1443 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1443              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1444 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
  1444   apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
  1445 apply(perm_simp)
  1445   apply(perm_simp)
  1446 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1446   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1447 apply(drule sym)
  1447   apply(drule sym)
  1448 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1448   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1449 apply(simp add: eqvts calc_atm)
  1449   apply(simp add: eqvts calc_atm)
  1450 done
  1450   done
  1451 
  1451 
  1452 lemma nrename_OrR2:
  1452 lemma nrename_OrR2:
  1453   assumes a: "R[x\<turnstile>n>y] = OrR2 <c>.N d" "c\<sharp>(R,d)" 
  1453   assumes a: "R[x\<turnstile>n>y] = OrR2 <c>.N d" "c\<sharp>(R,d)" 
  1454   shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
  1454   shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
  1455 using a
  1455   using a
  1456 apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
  1456   apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
  1457 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1457              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1458 apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
  1458   apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
  1459 apply(perm_simp)
  1459   apply(perm_simp)
  1460 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1460   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1461 apply(drule sym)
  1461   apply(drule sym)
  1462 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1462   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1463 apply(simp add: eqvts calc_atm)
  1463   apply(simp add: eqvts calc_atm)
  1464 done
  1464   done
  1465 
  1465 
  1466 lemma nrename_OrL:
  1466 lemma nrename_OrL:
  1467   assumes a: "R[u\<turnstile>n>v] = OrL (x).M (y).N z" "x\<sharp>(y,z,u,v,N,R)" "y\<sharp>(x,z,u,v,M,R)" "z\<sharp>(u,v)"
  1467   assumes a: "R[u\<turnstile>n>v] = OrL (x).M (y).N z" "x\<sharp>(y,z,u,v,N,R)" "y\<sharp>(x,z,u,v,M,R)" "z\<sharp>(u,v)"
  1468   shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[u\<turnstile>n>v] = M \<and> N'[u\<turnstile>n>v] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
  1468   shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[u\<turnstile>n>v] = M \<and> N'[u\<turnstile>n>v] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
  1469 using a
  1469   using a
  1470 apply(nominal_induct R avoiding: u v x y z M N rule: trm.strong_induct)
  1470   apply(nominal_induct R avoiding: u v x y z M N rule: trm.strong_induct)
  1471 apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
  1471              apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
  1472 apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
  1472   apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
  1473 apply(perm_simp)
  1473   apply(perm_simp)
  1474 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1474   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1475 apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
  1475   apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
  1476 apply(perm_simp)
  1476   apply(perm_simp)
  1477 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1477   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1478 apply(drule sym)
  1478    apply(drule sym)
  1479 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1479    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1480 apply(simp add: eqvts calc_atm)
  1480    apply(simp add: eqvts calc_atm)
  1481 apply(drule_tac s="trm2[u\<turnstile>n>v]" in  sym)
  1481   apply(drule_tac s="trm2[u\<turnstile>n>v]" in  sym)
  1482 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1482   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1483 apply(simp add: eqvts calc_atm)
  1483   apply(simp add: eqvts calc_atm)
  1484 done
  1484   done
  1485 
  1485 
  1486 lemma nrename_OrL':
  1486 lemma nrename_OrL':
  1487   assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" "v\<sharp>(R,N,u,x,y)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
  1487   assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" "v\<sharp>(R,N,u,x,y)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
  1488   shows "(\<exists>M' N'. (R = OrL (v).M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
  1488   shows "(\<exists>M' N'. (R = OrL (v).M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
  1489          (\<exists>M' N'. (R = OrL (v).M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
  1489          (\<exists>M' N'. (R = OrL (v).M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
  1490 using a [[simproc del: defined_all]]
  1490   using a [[simproc del: defined_all]]
  1491 apply(nominal_induct R avoiding: y x u v w M N rule: trm.strong_induct)
  1491   apply(nominal_induct R avoiding: y x u v w M N rule: trm.strong_induct)
  1492 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1492              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1493 apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
  1493    apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
  1494 apply(perm_simp)
  1494    apply(perm_simp)
  1495 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1495    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1496 apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
  1496    apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
  1497 apply(perm_simp)
  1497    apply(perm_simp)
  1498 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1498    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1499 apply(rule conjI)
  1499    apply(rule conjI)
  1500 apply(drule sym)
  1500     apply(drule sym)
  1501 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1501     apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1502 apply(simp add: eqvts calc_atm)
  1502     apply(simp add: eqvts calc_atm)
  1503 apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
  1503    apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
  1504 apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1504    apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1505 apply(simp add: eqvts calc_atm)
  1505    apply(simp add: eqvts calc_atm)
  1506 apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
  1506   apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
  1507 apply(perm_simp)
  1507   apply(perm_simp)
  1508 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1508   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1509 apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
  1509   apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
  1510 apply(perm_simp)
  1510   apply(perm_simp)
  1511 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1511   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1512 apply(rule conjI)
  1512   apply(rule conjI)
  1513 apply(drule sym)
  1513    apply(drule sym)
  1514 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1514    apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1515 apply(simp add: eqvts calc_atm)
  1515    apply(simp add: eqvts calc_atm)
  1516 apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
  1516   apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
  1517 apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1517   apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1518 apply(simp add: eqvts calc_atm)
  1518   apply(simp add: eqvts calc_atm)
  1519 done
  1519   done
  1520 
  1520 
  1521 lemma nrename_OrL_aux:
  1521 lemma nrename_OrL_aux:
  1522   assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" 
  1522   assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" 
  1523   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1523   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1524 using a
  1524   using a
  1525 apply(nominal_induct R avoiding: y x w u v M N rule: trm.strong_induct)
  1525   apply(nominal_induct R avoiding: y x w u v M N rule: trm.strong_induct)
  1526 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1526              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1527 done
  1527   done
  1528 
  1528 
  1529 lemma nrename_ImpL:
  1529 lemma nrename_ImpL:
  1530   assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (u).N z" "c\<sharp>(N,R)" "u\<sharp>(y,x,z,M,R)" "z\<sharp>(x,y)"
  1530   assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (u).N z" "c\<sharp>(N,R)" "u\<sharp>(y,x,z,M,R)" "z\<sharp>(x,y)"
  1531   shows "\<exists>M' N'. R = ImpL <c>.M' (u).N' z \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> u\<sharp>M'"
  1531   shows "\<exists>M' N'. R = ImpL <c>.M' (u).N' z \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> u\<sharp>M'"
  1532 using a
  1532   using a
  1533 apply(nominal_induct R avoiding: u x c y z M N rule: trm.strong_induct)
  1533   apply(nominal_induct R avoiding: u x c y z M N rule: trm.strong_induct)
  1534 apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
  1534              apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
  1535 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1535   apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1536 apply(perm_simp)
  1536   apply(perm_simp)
  1537 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1537   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1538 apply(rule_tac x="[(name1,u)]\<bullet>trm2" in exI)
  1538   apply(rule_tac x="[(name1,u)]\<bullet>trm2" in exI)
  1539 apply(perm_simp)
  1539   apply(perm_simp)
  1540 apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1540   apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
  1541 apply(drule sym)
  1541    apply(drule sym)
  1542 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1542    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1543 apply(simp add: eqvts calc_atm)
  1543    apply(simp add: eqvts calc_atm)
  1544 apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
  1544   apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
  1545 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1545   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1546 apply(simp add: eqvts calc_atm fresh_prod fresh_atm)
  1546   apply(simp add: eqvts calc_atm fresh_prod fresh_atm)
  1547 done
  1547   done
  1548 
  1548 
  1549 lemma nrename_ImpL':
  1549 lemma nrename_ImpL':
  1550   assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" "c\<sharp>(R,N)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
  1550   assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" "c\<sharp>(R,N)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
  1551   shows "(\<exists>M' N'. (R = ImpL <c>.M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
  1551   shows "(\<exists>M' N'. (R = ImpL <c>.M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
  1552          (\<exists>M' N'. (R = ImpL <c>.M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
  1552          (\<exists>M' N'. (R = ImpL <c>.M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
  1553 using a [[simproc del: defined_all]]
  1553   using a [[simproc del: defined_all]]
  1554 apply(nominal_induct R avoiding: y x u c w M N rule: trm.strong_induct)
  1554   apply(nominal_induct R avoiding: y x u c w M N rule: trm.strong_induct)
  1555 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1555              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
  1556 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1556    apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1557 apply(perm_simp)
  1557    apply(perm_simp)
  1558 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1558    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1559 apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
  1559    apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
  1560 apply(perm_simp)
  1560    apply(perm_simp)
  1561 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1561    apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1562 apply(rule conjI)
  1562    apply(rule conjI)
  1563 apply(drule sym)
  1563     apply(drule sym)
  1564 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1564     apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1565 apply(simp add: eqvts calc_atm)
  1565     apply(simp add: eqvts calc_atm)
  1566 apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
  1566    apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
  1567 apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1567    apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1568 apply(simp add: eqvts calc_atm)
  1568    apply(simp add: eqvts calc_atm)
  1569 apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1569   apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
  1570 apply(perm_simp)
  1570   apply(perm_simp)
  1571 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1571   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1572 apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
  1572   apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
  1573 apply(perm_simp)
  1573   apply(perm_simp)
  1574 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1574   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1575 apply(rule conjI)
  1575   apply(rule conjI)
  1576 apply(drule sym)
  1576    apply(drule sym)
  1577 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1577    apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1578 apply(simp add: eqvts calc_atm)
  1578    apply(simp add: eqvts calc_atm)
  1579 apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
  1579   apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
  1580 apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1580   apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
  1581 apply(simp add: eqvts calc_atm)
  1581   apply(simp add: eqvts calc_atm)
  1582 done
  1582   done
  1583 
  1583 
  1584 lemma nrename_ImpL_aux:
  1584 lemma nrename_ImpL_aux:
  1585   assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" 
  1585   assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" 
  1586   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1586   shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
  1587 using a
  1587   using a
  1588 apply(nominal_induct R avoiding: y x w u c M N rule: trm.strong_induct)
  1588   apply(nominal_induct R avoiding: y x w u c M N rule: trm.strong_induct)
  1589 apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1589              apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
  1590 done
  1590   done
  1591 
  1591 
  1592 lemma nrename_ImpR:
  1592 lemma nrename_ImpR:
  1593   assumes a: "R[u\<turnstile>n>v] = ImpR (x).<c>.N d" "c\<sharp>(R,d)" "x\<sharp>(R,u,v)" 
  1593   assumes a: "R[u\<turnstile>n>v] = ImpR (x).<c>.N d" "c\<sharp>(R,d)" "x\<sharp>(R,u,v)" 
  1594   shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[u\<turnstile>n>v] = N" 
  1594   shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[u\<turnstile>n>v] = N" 
  1595 using a
  1595   using a
  1596 apply(nominal_induct R avoiding: u v x c d N rule: trm.strong_induct)
  1596   apply(nominal_induct R avoiding: u v x c d N rule: trm.strong_induct)
  1597 apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
  1597              apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
  1598 apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
  1598    apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
  1599 apply(perm_simp)
  1599    apply(perm_simp)
  1600 apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1600   apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
  1601 apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
  1601   apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
  1602 apply(perm_simp)
  1602   apply(perm_simp)
  1603 apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
  1603   apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
  1604 apply(drule sym)
  1604   apply(drule sym)
  1605 apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1605   apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
  1606 apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1606   apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
  1607 apply(simp add: eqvts calc_atm)
  1607   apply(simp add: eqvts calc_atm)
  1608 done
  1608   done
  1609 
  1609 
  1610 lemma nrename_credu:
  1610 lemma nrename_credu:
  1611   assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>c M'"
  1611   assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>c M'"
  1612   shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>c M0"
  1612   shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>c M0"
  1613 using a
  1613   using a
  1614 apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: c_redu.strong_induct)
  1614   apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: c_redu.strong_induct)
  1615 apply(drule sym)
  1615    apply(drule sym)
  1616 apply(drule nrename_Cut)
  1616    apply(drule nrename_Cut)
  1617 apply(simp)
  1617      apply(simp)
  1618 apply(simp)
  1618     apply(simp)
  1619 apply(auto)
  1619    apply(auto)
  1620 apply(rule_tac x="M'{a:=(x).N'}" in exI)
  1620    apply(rule_tac x="M'{a:=(x).N'}" in exI)
  1621 apply(rule conjI)
  1621    apply(rule conjI)
  1622 apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
  1622     apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
  1623 apply(rule c_redu.intros)
  1623    apply(rule c_redu.intros)
  1624 apply(auto dest: not_fic_nrename)[1]
  1624      apply(auto dest: not_fic_nrename)[1]
  1625 apply(simp add: abs_fresh)
  1625     apply(simp add: abs_fresh)
  1626 apply(simp add: abs_fresh)
  1626    apply(simp add: abs_fresh)
  1627 apply(drule sym)
  1627   apply(drule sym)
  1628 apply(drule nrename_Cut)
  1628   apply(drule nrename_Cut)
  1629 apply(simp)
  1629     apply(simp)
  1630 apply(simp)
  1630    apply(simp)
  1631 apply(auto)
  1631   apply(auto)
  1632 apply(rule_tac x="N'{x:=<a>.M'}" in exI)
  1632   apply(rule_tac x="N'{x:=<a>.M'}" in exI)
  1633 apply(rule conjI)
  1633   apply(rule conjI)
  1634 apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
  1634    apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
  1635 apply(rule c_redu.intros)
  1635   apply(rule c_redu.intros)
  1636 apply(auto)
  1636     apply(auto)
  1637 apply(drule_tac x="xa" and y="y" in fin_nrename)
  1637   apply(drule_tac x="xa" and y="y" in fin_nrename)
  1638 apply(auto simp add: fresh_prod abs_fresh)
  1638    apply(auto simp add: fresh_prod abs_fresh)
  1639 done
  1639   done
  1640 
  1640 
  1641 lemma nrename_ax2:
  1641 lemma nrename_ax2:
  1642   assumes a: "N[x\<turnstile>n>y] = Ax z c"
  1642   assumes a: "N[x\<turnstile>n>y] = Ax z c"
  1643   shows "\<exists>z. N = Ax z c"
  1643   shows "\<exists>z. N = Ax z c"
  1644 using a
  1644   using a
  1645 apply(nominal_induct N avoiding: x y rule: trm.strong_induct)
  1645   apply(nominal_induct N avoiding: x y rule: trm.strong_induct)
  1646 apply(auto split: if_splits)
  1646              apply(auto split: if_splits)
  1647 apply(simp add: trm.inject)
  1647   apply(simp add: trm.inject)
  1648 done
  1648   done
  1649 
  1649 
  1650 lemma fic_nrename:
  1650 lemma fic_nrename:
  1651   assumes a: "fic (M[x\<turnstile>n>y]) c" 
  1651   assumes a: "fic (M[x\<turnstile>n>y]) c" 
  1652   shows "fic M c" 
  1652   shows "fic M c" 
  1653 using a
  1653   using a
  1654 apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
  1654   apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
  1655 apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
  1655              apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
  1656            split: if_splits)
  1656       split: if_splits)
  1657 apply(auto dest: nrename_fresh_interesting2 simp add: fresh_prod fresh_atm)
  1657        apply(auto dest: nrename_fresh_interesting2 simp add: fresh_prod fresh_atm)
  1658 done
  1658   done
  1659 
  1659 
  1660 lemma nrename_lredu:
  1660 lemma nrename_lredu:
  1661   assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>l M'"
  1661   assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>l M'"
  1662   shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>l M0"
  1662   shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>l M0"
  1663 using a
  1663   using a
  1664 apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: l_redu.strong_induct)
  1664   apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: l_redu.strong_induct)
  1665 apply(drule sym)
  1665          apply(drule sym)
  1666 apply(drule nrename_Cut)
  1666          apply(drule nrename_Cut)
  1667 apply(simp add: fresh_prod fresh_atm)
  1667            apply(simp add: fresh_prod fresh_atm)
  1668 apply(simp)
  1668           apply(simp)
  1669 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1669          apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1670 apply(case_tac "xa=y")
  1670          apply(case_tac "xa=y")
  1671 apply(simp add: nrename_id)
  1671           apply(simp add: nrename_id)
  1672 apply(rule l_redu.intros)
  1672           apply(rule l_redu.intros)
  1673 apply(simp)
  1673             apply(simp)
  1674 apply(simp add: fresh_atm)
  1674            apply(simp add: fresh_atm)
  1675 apply(assumption)
  1675           apply(assumption)
  1676 apply(frule nrename_ax2)
  1676          apply(frule nrename_ax2)
  1677 apply(auto)[1]
  1677          apply(auto)[1]
  1678 apply(case_tac "z=xa")
  1678          apply(case_tac "z=xa")
  1679 apply(simp add: trm.inject)
  1679           apply(simp add: trm.inject)
  1680 apply(simp)
  1680          apply(simp)
  1681 apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
  1681          apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
  1682 apply(rule conjI)
  1682          apply(rule conjI)
  1683 apply(rule crename_interesting3)
  1683           apply(rule crename_interesting3)
  1684 apply(rule l_redu.intros)
  1684          apply(rule l_redu.intros)
  1685 apply(simp)
  1685            apply(simp)
  1686 apply(simp add: fresh_atm)
  1686           apply(simp add: fresh_atm)
  1687 apply(auto dest: fic_nrename simp add: fresh_prod fresh_atm)[1]
  1687          apply(auto dest: fic_nrename simp add: fresh_prod fresh_atm)[1]
  1688 apply(drule sym)
  1688         apply(drule sym)
  1689 apply(drule nrename_Cut)
  1689         apply(drule nrename_Cut)
  1690 apply(simp add: fresh_prod fresh_atm)
  1690           apply(simp add: fresh_prod fresh_atm)
  1691 apply(simp add: fresh_prod fresh_atm)
  1691          apply(simp add: fresh_prod fresh_atm)
  1692 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1692         apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1693 apply(case_tac "xa=ya")
  1693         apply(case_tac "xa=ya")
  1694 apply(simp add: nrename_id)
  1694          apply(simp add: nrename_id)
  1695 apply(rule l_redu.intros)
  1695          apply(rule l_redu.intros)
  1696 apply(simp)
  1696            apply(simp)
  1697 apply(simp add: fresh_atm)
  1697           apply(simp add: fresh_atm)
  1698 apply(assumption)
  1698          apply(assumption)
  1699 apply(frule nrename_ax2)
  1699         apply(frule nrename_ax2)
  1700 apply(auto)[1]
  1700         apply(auto)[1]
  1701 apply(case_tac "z=xa")
  1701         apply(case_tac "z=xa")
  1702 apply(simp add: trm.inject)
  1702          apply(simp add: trm.inject)
  1703 apply(rule_tac x="N'[x\<turnstile>n>xa]" in exI)
  1703          apply(rule_tac x="N'[x\<turnstile>n>xa]" in exI)
  1704 apply(rule conjI)
  1704          apply(rule conjI)
  1705 apply(rule nrename_interesting1)
  1705           apply(rule nrename_interesting1)
  1706 apply(auto)[1]
  1706           apply(auto)[1]
  1707 apply(rule l_redu.intros)
  1707          apply(rule l_redu.intros)
  1708 apply(simp)
  1708            apply(simp)
  1709 apply(simp add: fresh_atm)
  1709           apply(simp add: fresh_atm)
  1710 apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
  1710          apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
  1711 apply(simp add: trm.inject)
  1711         apply(simp add: trm.inject)
  1712 apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
  1712         apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
  1713 apply(rule conjI)
  1713         apply(rule conjI)
  1714 apply(rule nrename_interesting2)
  1714          apply(rule nrename_interesting2)
  1715 apply(simp_all)
  1715              apply(simp_all)
  1716 apply(rule l_redu.intros)
  1716         apply(rule l_redu.intros)
  1717 apply(simp)
  1717           apply(simp)
  1718 apply(simp add: fresh_atm)
  1718          apply(simp add: fresh_atm)
  1719 apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
  1719         apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
  1720 (* LNot *)
  1720     (* LNot *)
  1721 apply(drule sym)
  1721        apply(drule sym)
  1722 apply(drule nrename_Cut)
  1722        apply(drule nrename_Cut)
  1723 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1723          apply(simp add: fresh_prod abs_fresh fresh_atm)
  1724 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1724         apply(simp add: fresh_prod abs_fresh fresh_atm)
  1725 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1725        apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1726 apply(drule nrename_NotR)
  1726        apply(drule nrename_NotR)
  1727 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1727         apply(simp add: fresh_prod abs_fresh fresh_atm)
  1728 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1728        apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1729 apply(drule nrename_NotL)
  1729        apply(drule nrename_NotL)
  1730 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1730          apply(simp add: fresh_prod abs_fresh fresh_atm)
  1731 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1731         apply(simp add: fresh_prod abs_fresh fresh_atm)
  1732 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1732        apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1733 apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
  1733        apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
  1734 apply(simp add: fresh_atm)[1]
  1734        apply(simp add: fresh_atm)[1]
  1735 apply(rule l_redu.intros)
  1735        apply(rule l_redu.intros)
  1736 apply(auto simp add: fresh_prod fresh_atm intro: nrename_fresh_interesting1)[1]
  1736             apply(auto simp add: fresh_prod fresh_atm intro: nrename_fresh_interesting1)[1]
  1737 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1737            apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1738 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
  1738           apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
  1739 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
  1739          apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
  1740 apply(simp add: fresh_atm)
  1740         apply(simp add: fresh_atm)
  1741 apply(simp add: fresh_atm)
  1741        apply(simp add: fresh_atm)
  1742 (* LAnd1 *)
  1742     (* LAnd1 *)
  1743 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1743       apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1744 apply(drule sym)
  1744       apply(drule sym)
  1745 apply(drule nrename_Cut)
  1745       apply(drule nrename_Cut)
  1746 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1746         apply(simp add: fresh_prod abs_fresh fresh_atm)
  1747 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1747        apply(simp add: fresh_prod abs_fresh fresh_atm)
  1748 apply(auto)[1]
  1748       apply(auto)[1]
  1749 apply(drule nrename_AndR)
  1749       apply(drule nrename_AndR)
  1750 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1750         apply(simp add: fresh_prod abs_fresh fresh_atm)
  1751 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1751        apply(simp add: fresh_prod abs_fresh fresh_atm)
  1752 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1752       apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1753 apply(drule nrename_AndL1)
  1753       apply(drule nrename_AndL1)
  1754 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1754         apply(simp add: fresh_prod abs_fresh fresh_atm)
  1755 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1755        apply(simp add: fresh_prod abs_fresh fresh_atm)
  1756 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1756       apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1757 apply(rule_tac x="Cut <a1>.M'a (x).N'b" in exI)
  1757       apply(rule_tac x="Cut <a1>.M'a (x).N'b" in exI)
  1758 apply(simp add: fresh_atm)[1]
  1758       apply(simp add: fresh_atm)[1]
  1759 apply(rule l_redu.intros)
  1759       apply(rule l_redu.intros)
  1760 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1760            apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1761 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1761           apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1762 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1762          apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1763 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1763         apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1764 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1764        apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1765 apply(simp add: fresh_atm)
  1765       apply(simp add: fresh_atm)
  1766 (* LAnd2 *)
  1766     (* LAnd2 *)
  1767 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1767      apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1768 apply(drule sym)
  1768      apply(drule sym)
  1769 apply(drule nrename_Cut)
  1769      apply(drule nrename_Cut)
  1770 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1770        apply(simp add: fresh_prod abs_fresh fresh_atm)
  1771 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1771       apply(simp add: fresh_prod abs_fresh fresh_atm)
  1772 apply(auto)[1]
  1772      apply(auto)[1]
  1773 apply(drule nrename_AndR)
  1773      apply(drule nrename_AndR)
  1774 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1774        apply(simp add: fresh_prod abs_fresh fresh_atm)
  1775 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1775       apply(simp add: fresh_prod abs_fresh fresh_atm)
  1776 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1776      apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1777 apply(drule nrename_AndL2)
  1777      apply(drule nrename_AndL2)
  1778 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1778        apply(simp add: fresh_prod abs_fresh fresh_atm)
  1779 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1779       apply(simp add: fresh_prod abs_fresh fresh_atm)
  1780 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1780      apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1781 apply(rule_tac x="Cut <a2>.N'a (x).N'b" in exI)
  1781      apply(rule_tac x="Cut <a2>.N'a (x).N'b" in exI)
  1782 apply(simp add: fresh_atm)[1]
  1782      apply(simp add: fresh_atm)[1]
  1783 apply(rule l_redu.intros)
  1783      apply(rule l_redu.intros)
  1784 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1784           apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1785 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1785          apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1786 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1786         apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1787 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1787        apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1788 apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1788       apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1789 apply(simp add: fresh_atm)
  1789      apply(simp add: fresh_atm)
  1790 (* LOr1 *)
  1790     (* LOr1 *)
  1791 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1791     apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1792 apply(drule sym)
  1792     apply(drule sym)
  1793 apply(drule nrename_Cut)
  1793     apply(drule nrename_Cut)
  1794 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1794       apply(simp add: fresh_prod abs_fresh fresh_atm)
  1795 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1795      apply(simp add: fresh_prod abs_fresh fresh_atm)
  1796 apply(auto)[1]
  1796     apply(auto)[1]
  1797 apply(drule nrename_OrL)
  1797     apply(drule nrename_OrL)
  1798 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1798        apply(simp add: fresh_prod abs_fresh fresh_atm)
  1799 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1799       apply(simp add: fresh_prod abs_fresh fresh_atm)
  1800 apply(simp add: fresh_prod fresh_atm)
  1800      apply(simp add: fresh_prod fresh_atm)
  1801 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1801     apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1802 apply(drule nrename_OrR1)
  1802     apply(drule nrename_OrR1)
  1803 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1803      apply(simp add: fresh_prod abs_fresh fresh_atm)
  1804 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1804     apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1805 apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
  1805     apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
  1806 apply(rule conjI)
  1806     apply(rule conjI)
  1807 apply(simp add: abs_fresh fresh_atm)[1]
  1807      apply(simp add: abs_fresh fresh_atm)[1]
  1808 apply(rule l_redu.intros)
  1808     apply(rule l_redu.intros)
  1809 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1809          apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1810 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1810         apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1811 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1811        apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1812 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1812       apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1813 apply(simp add: fresh_atm)
  1813      apply(simp add: fresh_atm)
  1814 apply(simp add: fresh_atm)
  1814     apply(simp add: fresh_atm)
  1815 (* LOr2 *)
  1815     (* LOr2 *)
  1816 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1816    apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1817 apply(drule sym)
  1817    apply(drule sym)
  1818 apply(drule nrename_Cut)
  1818    apply(drule nrename_Cut)
  1819 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1819      apply(simp add: fresh_prod abs_fresh fresh_atm)
  1820 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1820     apply(simp add: fresh_prod abs_fresh fresh_atm)
  1821 apply(auto)[1]
  1821    apply(auto)[1]
  1822 apply(drule nrename_OrL)
  1822    apply(drule nrename_OrL)
  1823 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1823       apply(simp add: fresh_prod abs_fresh fresh_atm)
  1824 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1824      apply(simp add: fresh_prod abs_fresh fresh_atm)
  1825 apply(simp add: fresh_prod fresh_atm)
  1825     apply(simp add: fresh_prod fresh_atm)
  1826 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1826    apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1827 apply(drule nrename_OrR2)
  1827    apply(drule nrename_OrR2)
  1828 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1828     apply(simp add: fresh_prod abs_fresh fresh_atm)
  1829 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1829    apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1830 apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
  1830    apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
  1831 apply(rule conjI)
  1831    apply(rule conjI)
  1832 apply(simp add: abs_fresh fresh_atm)[1]
  1832     apply(simp add: abs_fresh fresh_atm)[1]
  1833 apply(rule l_redu.intros)
  1833    apply(rule l_redu.intros)
  1834 apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1834         apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
  1835 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1835        apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1836 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1836       apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1837 apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1837      apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
  1838 apply(simp add: fresh_atm)
  1838     apply(simp add: fresh_atm)
  1839 apply(simp add: fresh_atm)
  1839    apply(simp add: fresh_atm)
  1840 (* ImpL *)
  1840     (* ImpL *)
  1841 apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1841   apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
  1842 apply(drule sym) 
  1842   apply(drule sym) 
  1843 apply(drule nrename_Cut)
  1843   apply(drule nrename_Cut)
  1844 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1844     apply(simp add: fresh_prod abs_fresh fresh_atm)
  1845 apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
  1845    apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
  1846 apply(auto)[1]
  1846   apply(auto)[1]
  1847 apply(drule nrename_ImpL)
  1847   apply(drule nrename_ImpL)
  1848 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1848      apply(simp add: fresh_prod abs_fresh fresh_atm)
  1849 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1849     apply(simp add: fresh_prod abs_fresh fresh_atm)
  1850 apply(simp add: fresh_prod fresh_atm)
  1850    apply(simp add: fresh_prod fresh_atm)
  1851 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1851   apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1852 apply(drule nrename_ImpR)
  1852   apply(drule nrename_ImpR)
  1853 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1853     apply(simp add: fresh_prod abs_fresh fresh_atm)
  1854 apply(simp add: fresh_prod abs_fresh fresh_atm)
  1854    apply(simp add: fresh_prod abs_fresh fresh_atm)
  1855 apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1855   apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
  1856 apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
  1856   apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
  1857 apply(rule conjI)
  1857   apply(rule conjI)
  1858 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1858    apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1859 apply(rule l_redu.intros)
  1859   apply(rule l_redu.intros)
  1860 apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
  1860        apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
  1861 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1861       apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1862 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
  1862      apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
  1863 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1863     apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1864 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1864    apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1865 apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1865   apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
  1866 done
  1866   done
  1867 
  1867 
  1868 lemma nrename_aredu:
  1868 lemma nrename_aredu:
  1869   assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>a M'" "x\<noteq>y"
  1869   assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>a M'" "x\<noteq>y"
  1870   shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>a M0"
  1870   shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>a M0"
  1871 using a
  1871   using a
  1872 apply(nominal_induct "M[x\<turnstile>n>y]" M' avoiding: M x y rule: a_redu.strong_induct)
  1872   apply(nominal_induct "M[x\<turnstile>n>y]" M' avoiding: M x y rule: a_redu.strong_induct)
  1873 apply(drule  nrename_lredu)
  1873                   apply(drule  nrename_lredu)
  1874 apply(blast)
  1874                   apply(blast)
  1875 apply(drule  nrename_credu)
  1875                  apply(drule  nrename_credu)
  1876 apply(blast)
  1876                  apply(blast)
  1877 (* Cut *)
  1877     (* Cut *)
  1878 apply(drule sym)
  1878                 apply(drule sym)
  1879 apply(drule nrename_Cut)
  1879                 apply(drule nrename_Cut)
  1880 apply(simp)
  1880                   apply(simp)
  1881 apply(simp)
  1881                  apply(simp)
  1882 apply(auto)[1]
  1882                 apply(auto)[1]
  1883 apply(drule_tac x="M'a" in meta_spec)
  1883                 apply(drule_tac x="M'a" in meta_spec)
  1884 apply(drule_tac x="xa" in meta_spec)
  1884                 apply(drule_tac x="xa" in meta_spec)
  1885 apply(drule_tac x="y" in meta_spec)
  1885                 apply(drule_tac x="y" in meta_spec)
  1886 apply(auto)[1]
  1886                 apply(auto)[1]
  1887 apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
  1887                 apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
  1888 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1888                 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1889 apply(rule conjI)
  1889                 apply(rule conjI)
  1890 apply(rule trans)
  1890                  apply(rule trans)
  1891 apply(rule nrename.simps)
  1891                   apply(rule nrename.simps)
  1892 apply(drule nrename_fresh_interesting2)
  1892                    apply(drule nrename_fresh_interesting2)
  1893 apply(simp add: fresh_a_redu)
  1893                    apply(simp add: fresh_a_redu)
  1894 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1894                   apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1895 apply(drule nrename_fresh_interesting1)
  1895                   apply(drule nrename_fresh_interesting1)
  1896 apply(simp add: fresh_prod fresh_atm)
  1896                    apply(simp add: fresh_prod fresh_atm)
  1897 apply(simp add: fresh_a_redu)
  1897                   apply(simp add: fresh_a_redu)
  1898 apply(simp)
  1898                  apply(simp)
  1899 apply(auto)[1]
  1899                 apply(auto)[1]
  1900 apply(drule sym)
  1900                apply(drule sym)
  1901 apply(drule nrename_Cut)
  1901                apply(drule nrename_Cut)
  1902 apply(simp)
  1902                  apply(simp)
  1903 apply(simp)
  1903                 apply(simp)
  1904 apply(auto)[1]
  1904                apply(auto)[1]
  1905 apply(drule_tac x="N'a" in meta_spec)
  1905                apply(drule_tac x="N'a" in meta_spec)
  1906 apply(drule_tac x="xa" in meta_spec)
  1906                apply(drule_tac x="xa" in meta_spec)
  1907 apply(drule_tac x="y" in meta_spec)
  1907                apply(drule_tac x="y" in meta_spec)
  1908 apply(auto)[1]
  1908                apply(auto)[1]
  1909 apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
  1909                apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
  1910 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1910                apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1911 apply(rule conjI)
  1911                apply(rule conjI)
  1912 apply(rule trans)
  1912                 apply(rule trans)
  1913 apply(rule nrename.simps)
  1913                  apply(rule nrename.simps)
  1914 apply(simp add: fresh_a_redu)
  1914                   apply(simp add: fresh_a_redu)
  1915 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1915                  apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1916 apply(simp)
  1916                 apply(simp)
  1917 apply(auto)[1]
  1917                apply(auto)[1]
  1918 (* NotL *)
  1918     (* NotL *)
  1919 apply(drule sym)
  1919               apply(drule sym)
  1920 apply(frule nrename_NotL_aux)
  1920               apply(frule nrename_NotL_aux)
  1921 apply(erule disjE)
  1921               apply(erule disjE)
  1922 apply(auto)[1]
  1922                apply(auto)[1]
  1923 apply(drule nrename_NotL')
  1923               apply(drule nrename_NotL')
  1924 apply(simp)
  1924                 apply(simp)
  1925 apply(simp add: fresh_atm)
  1925                apply(simp add: fresh_atm)
  1926 apply(erule disjE)
  1926               apply(erule disjE)
  1927 apply(auto)[1]
  1927                apply(auto)[1]
  1928 apply(drule_tac x="N'" in meta_spec)
  1928                apply(drule_tac x="N'" in meta_spec)
  1929 apply(drule_tac x="xa" in meta_spec)
  1929                apply(drule_tac x="xa" in meta_spec)
  1930 apply(drule_tac x="y" in meta_spec)
  1930                apply(drule_tac x="y" in meta_spec)
  1931 apply(auto)[1]
  1931                apply(auto)[1]
  1932 apply(rule_tac x="NotL <a>.M0 x" in exI)
  1932                apply(rule_tac x="NotL <a>.M0 x" in exI)
  1933 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1933                apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1934 apply(auto)[1]
  1934                apply(auto)[1]
  1935 apply(auto)[1]
  1935               apply(auto)[1]
  1936 apply(drule_tac x="N'" in meta_spec)
  1936               apply(drule_tac x="N'" in meta_spec)
  1937 apply(drule_tac x="xa" in meta_spec)
  1937               apply(drule_tac x="xa" in meta_spec)
  1938 apply(drule_tac x="x" in meta_spec)
  1938               apply(drule_tac x="x" in meta_spec)
  1939 apply(auto)[1]
  1939               apply(auto)[1]
  1940 apply(rule_tac x="NotL <a>.M0 xa" in exI)
  1940               apply(rule_tac x="NotL <a>.M0 xa" in exI)
  1941 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1941               apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1942 apply(auto)[1]
  1942               apply(auto)[1]
  1943 (* NotR *)
  1943     (* NotR *)
  1944 apply(drule sym)
  1944              apply(drule sym)
  1945 apply(drule nrename_NotR)
  1945              apply(drule nrename_NotR)
  1946 apply(simp)
  1946               apply(simp)
  1947 apply(auto)[1]
  1947              apply(auto)[1]
  1948 apply(drule_tac x="N'" in meta_spec)
  1948              apply(drule_tac x="N'" in meta_spec)
  1949 apply(drule_tac x="xa" in meta_spec)
  1949              apply(drule_tac x="xa" in meta_spec)
  1950 apply(drule_tac x="y" in meta_spec)
  1950              apply(drule_tac x="y" in meta_spec)
  1951 apply(auto)[1]
  1951              apply(auto)[1]
  1952 apply(rule_tac x="NotR (x).M0 a" in exI)
  1952              apply(rule_tac x="NotR (x).M0 a" in exI)
  1953 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1953              apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1954 apply(auto)[1]
  1954              apply(auto)[1]
  1955 (* AndR *)
  1955     (* AndR *)
  1956 apply(drule sym)
  1956             apply(drule sym)
  1957 apply(drule nrename_AndR)
  1957             apply(drule nrename_AndR)
  1958 apply(simp)
  1958               apply(simp)
  1959 apply(auto simp add: fresh_atm fresh_prod)[1]
  1959               apply(auto simp add: fresh_atm fresh_prod)[1]
  1960 apply(auto simp add: fresh_atm fresh_prod)[1]
  1960              apply(auto simp add: fresh_atm fresh_prod)[1]
  1961 apply(auto)[1]
  1961             apply(auto)[1]
  1962 apply(drule_tac x="M'a" in meta_spec)
  1962             apply(drule_tac x="M'a" in meta_spec)
  1963 apply(drule_tac x="x" in meta_spec)
  1963             apply(drule_tac x="x" in meta_spec)
  1964 apply(drule_tac x="y" in meta_spec)
  1964             apply(drule_tac x="y" in meta_spec)
  1965 apply(auto)[1]
  1965             apply(auto)[1]
  1966 apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
  1966             apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
  1967 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1967             apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1968 apply(auto)[1]
  1968             apply(auto)[1]
  1969 apply(rule trans)
  1969             apply(rule trans)
  1970 apply(rule nrename.simps)
  1970              apply(rule nrename.simps)
  1971 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1971                apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1972 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1972               apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1973 apply(auto intro: fresh_a_redu)[1]
  1973               apply(auto intro: fresh_a_redu)[1]
  1974 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1974              apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1975 apply(simp)
  1975             apply(simp)
  1976 apply(drule sym)
  1976            apply(drule sym)
  1977 apply(drule nrename_AndR)
  1977            apply(drule nrename_AndR)
  1978 apply(simp)
  1978              apply(simp)
  1979 apply(auto simp add: fresh_atm fresh_prod)[1]
  1979              apply(auto simp add: fresh_atm fresh_prod)[1]
  1980 apply(auto simp add: fresh_atm fresh_prod)[1]
  1980             apply(auto simp add: fresh_atm fresh_prod)[1]
  1981 apply(auto)[1]
  1981            apply(auto)[1]
  1982 apply(drule_tac x="N'a" in meta_spec)
  1982            apply(drule_tac x="N'a" in meta_spec)
  1983 apply(drule_tac x="x" in meta_spec)
  1983            apply(drule_tac x="x" in meta_spec)
  1984 apply(drule_tac x="y" in meta_spec)
  1984            apply(drule_tac x="y" in meta_spec)
  1985 apply(auto)[1]
  1985            apply(auto)[1]
  1986 apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
  1986            apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
  1987 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1987            apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  1988 apply(auto)[1]
  1988            apply(auto)[1]
  1989 apply(rule trans)
  1989            apply(rule trans)
  1990 apply(rule nrename.simps)
  1990             apply(rule nrename.simps)
  1991 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1991               apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  1992 apply(auto intro: fresh_a_redu)[1]
  1992               apply(auto intro: fresh_a_redu)[1]
  1993 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1993              apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  1994 apply(simp)
  1994             apply(simp)
  1995 apply(simp)
  1995            apply(simp)
  1996 (* AndL1 *)
  1996     (* AndL1 *)
  1997 apply(drule sym)
  1997           apply(drule sym)
  1998 apply(frule nrename_AndL1_aux)
  1998           apply(frule nrename_AndL1_aux)
  1999 apply(erule disjE)
  1999           apply(erule disjE)
  2000 apply(auto)[1]
  2000            apply(auto)[1]
  2001 apply(drule nrename_AndL1')
  2001           apply(drule nrename_AndL1')
  2002 apply(simp)
  2002             apply(simp)
  2003 apply(simp add: fresh_atm)
  2003            apply(simp add: fresh_atm)
  2004 apply(erule disjE)
  2004           apply(erule disjE)
  2005 apply(auto)[1]
  2005            apply(auto)[1]
  2006 apply(drule_tac x="N'" in meta_spec)
  2006            apply(drule_tac x="N'" in meta_spec)
  2007 apply(drule_tac x="xa" in meta_spec)
  2007            apply(drule_tac x="xa" in meta_spec)
  2008 apply(drule_tac x="ya" in meta_spec)
  2008            apply(drule_tac x="ya" in meta_spec)
  2009 apply(auto)[1]
  2009            apply(auto)[1]
  2010 apply(rule_tac x="AndL1 (x).M0 y" in exI)
  2010            apply(rule_tac x="AndL1 (x).M0 y" in exI)
  2011 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2011            apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2012 apply(auto)[1]
  2012            apply(auto)[1]
  2013 apply(auto)[1]
  2013           apply(auto)[1]
  2014 apply(drule_tac x="N'" in meta_spec)
  2014           apply(drule_tac x="N'" in meta_spec)
  2015 apply(drule_tac x="xa" in meta_spec)
  2015           apply(drule_tac x="xa" in meta_spec)
  2016 apply(drule_tac x="y" in meta_spec)
  2016           apply(drule_tac x="y" in meta_spec)
  2017 apply(auto)[1]
  2017           apply(auto)[1]
  2018 apply(rule_tac x="AndL1 (x).M0 xa" in exI)
  2018           apply(rule_tac x="AndL1 (x).M0 xa" in exI)
  2019 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2019           apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2020 apply(auto)[1]
  2020           apply(auto)[1]
  2021 (* AndL2 *)
  2021     (* AndL2 *)
  2022 apply(drule sym)
  2022          apply(drule sym)
  2023 apply(frule nrename_AndL2_aux)
  2023          apply(frule nrename_AndL2_aux)
  2024 apply(erule disjE)
  2024          apply(erule disjE)
  2025 apply(auto)[1]
  2025           apply(auto)[1]
  2026 apply(drule nrename_AndL2')
  2026          apply(drule nrename_AndL2')
  2027 apply(simp)
  2027            apply(simp)
  2028 apply(simp add: fresh_atm)
  2028           apply(simp add: fresh_atm)
  2029 apply(erule disjE)
  2029          apply(erule disjE)
  2030 apply(auto)[1]
  2030           apply(auto)[1]
  2031 apply(drule_tac x="N'" in meta_spec)
  2031           apply(drule_tac x="N'" in meta_spec)
  2032 apply(drule_tac x="xa" in meta_spec)
  2032           apply(drule_tac x="xa" in meta_spec)
  2033 apply(drule_tac x="ya" in meta_spec)
  2033           apply(drule_tac x="ya" in meta_spec)
  2034 apply(auto)[1]
  2034           apply(auto)[1]
  2035 apply(rule_tac x="AndL2 (x).M0 y" in exI)
  2035           apply(rule_tac x="AndL2 (x).M0 y" in exI)
  2036 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2036           apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2037 apply(auto)[1]
  2037           apply(auto)[1]
  2038 apply(auto)[1]
  2038          apply(auto)[1]
  2039 apply(drule_tac x="N'" in meta_spec)
  2039          apply(drule_tac x="N'" in meta_spec)
  2040 apply(drule_tac x="xa" in meta_spec)
  2040          apply(drule_tac x="xa" in meta_spec)
  2041 apply(drule_tac x="y" in meta_spec)
  2041          apply(drule_tac x="y" in meta_spec)
  2042 apply(auto)[1]
  2042          apply(auto)[1]
  2043 apply(rule_tac x="AndL2 (x).M0 xa" in exI)
  2043          apply(rule_tac x="AndL2 (x).M0 xa" in exI)
  2044 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2044          apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2045 apply(auto)[1]
  2045          apply(auto)[1]
  2046 (* OrL *)
  2046     (* OrL *)
  2047 apply(drule sym)
  2047         apply(drule sym)
  2048 apply(frule nrename_OrL_aux)
  2048         apply(frule nrename_OrL_aux)
  2049 apply(erule disjE)
  2049         apply(erule disjE)
  2050 apply(auto)[1]
  2050          apply(auto)[1]
  2051 apply(drule nrename_OrL')
  2051         apply(drule nrename_OrL')
  2052 apply(simp add: fresh_prod fresh_atm)
  2052            apply(simp add: fresh_prod fresh_atm)
  2053 apply(simp add: fresh_atm)
  2053           apply(simp add: fresh_atm)
  2054 apply(simp add: fresh_atm)
  2054          apply(simp add: fresh_atm)
  2055 apply(erule disjE)
  2055         apply(erule disjE)
  2056 apply(auto)[1]
  2056          apply(auto)[1]
  2057 apply(drule_tac x="M'a" in meta_spec)
  2057          apply(drule_tac x="M'a" in meta_spec)
  2058 apply(drule_tac x="xa" in meta_spec)
  2058          apply(drule_tac x="xa" in meta_spec)
  2059 apply(drule_tac x="ya" in meta_spec)
  2059          apply(drule_tac x="ya" in meta_spec)
  2060 apply(auto)[1]
  2060          apply(auto)[1]
  2061 apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
  2061          apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
  2062 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2062          apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2063 apply(auto)[1]
  2063          apply(auto)[1]
  2064 apply(rule trans)
  2064          apply(rule trans)
  2065 apply(rule nrename.simps)
  2065           apply(rule nrename.simps)
  2066 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2066             apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2067 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2067            apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2068 apply(auto intro: fresh_a_redu)[1]
  2068            apply(auto intro: fresh_a_redu)[1]
  2069 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2069           apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2070 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2070          apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2071 apply(auto)[1]
  2071         apply(auto)[1]
  2072 apply(drule_tac x="M'a" in meta_spec)
  2072         apply(drule_tac x="M'a" in meta_spec)
  2073 apply(drule_tac x="xa" in meta_spec)
  2073         apply(drule_tac x="xa" in meta_spec)
  2074 apply(drule_tac x="z" in meta_spec)
  2074         apply(drule_tac x="z" in meta_spec)
  2075 apply(auto)[1]
  2075         apply(auto)[1]
  2076 apply(rule_tac x="OrL (x).M0 (y).N' xa" in exI)
  2076         apply(rule_tac x="OrL (x).M0 (y).N' xa" in exI)
  2077 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2077         apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2078 apply(auto)[1]
  2078         apply(auto)[1]
  2079 apply(rule trans)
  2079         apply(rule trans)
  2080 apply(rule nrename.simps)
  2080          apply(rule nrename.simps)
  2081 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2081            apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2082 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2082           apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2083 apply(auto intro: fresh_a_redu)[1]
  2083           apply(auto intro: fresh_a_redu)[1]
  2084 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2084          apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2085 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2085         apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2086 apply(drule sym)
  2086        apply(drule sym)
  2087 apply(frule nrename_OrL_aux)
  2087        apply(frule nrename_OrL_aux)
  2088 apply(erule disjE)
  2088        apply(erule disjE)
  2089 apply(auto)[1]
  2089         apply(auto)[1]
  2090 apply(drule nrename_OrL')
  2090        apply(drule nrename_OrL')
  2091 apply(simp add: fresh_prod fresh_atm)
  2091           apply(simp add: fresh_prod fresh_atm)
  2092 apply(simp add: fresh_atm)
  2092          apply(simp add: fresh_atm)
  2093 apply(simp add: fresh_atm)
  2093         apply(simp add: fresh_atm)
  2094 apply(erule disjE)
  2094        apply(erule disjE)
  2095 apply(auto)[1]
  2095         apply(auto)[1]
  2096 apply(drule_tac x="N'a" in meta_spec)
  2096         apply(drule_tac x="N'a" in meta_spec)
  2097 apply(drule_tac x="xa" in meta_spec)
  2097         apply(drule_tac x="xa" in meta_spec)
  2098 apply(drule_tac x="ya" in meta_spec)
  2098         apply(drule_tac x="ya" in meta_spec)
  2099 apply(auto)[1]
  2099         apply(auto)[1]
  2100 apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
  2100         apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
  2101 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2101         apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2102 apply(auto)[1]
  2102         apply(auto)[1]
  2103 apply(rule trans)
  2103         apply(rule trans)
  2104 apply(rule nrename.simps)
  2104          apply(rule nrename.simps)
  2105 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  2105            apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  2106 apply(auto intro: fresh_a_redu)[1]
  2106            apply(auto intro: fresh_a_redu)[1]
  2107 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2107           apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2108 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2108          apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2109 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2109         apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2110 apply(auto)[1]
  2110        apply(auto)[1]
  2111 apply(drule_tac x="N'a" in meta_spec)
  2111        apply(drule_tac x="N'a" in meta_spec)
  2112 apply(drule_tac x="xa" in meta_spec)
  2112        apply(drule_tac x="xa" in meta_spec)
  2113 apply(drule_tac x="z" in meta_spec)
  2113        apply(drule_tac x="z" in meta_spec)
  2114 apply(auto)[1]
  2114        apply(auto)[1]
  2115 apply(rule_tac x="OrL (x).M' (y).M0 xa" in exI)
  2115        apply(rule_tac x="OrL (x).M' (y).M0 xa" in exI)
  2116 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2116        apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2117 apply(auto)[1]
  2117        apply(auto)[1]
  2118 apply(rule trans)
  2118        apply(rule trans)
  2119 apply(rule nrename.simps)
  2119         apply(rule nrename.simps)
  2120 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  2120           apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
  2121 apply(auto intro: fresh_a_redu)[1]
  2121           apply(auto intro: fresh_a_redu)[1]
  2122 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2122          apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2123 apply(simp)
  2123         apply(simp)
  2124 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2124        apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2125 (* OrR1 *)
  2125     (* OrR1 *)
  2126 apply(drule sym)
  2126       apply(drule sym)
  2127 apply(drule nrename_OrR1)
  2127       apply(drule nrename_OrR1)
  2128 apply(simp)
  2128        apply(simp)
  2129 apply(auto)[1]
  2129       apply(auto)[1]
  2130 apply(drule_tac x="N'" in meta_spec)
  2130       apply(drule_tac x="N'" in meta_spec)
  2131 apply(drule_tac x="x" in meta_spec)
  2131       apply(drule_tac x="x" in meta_spec)
  2132 apply(drule_tac x="y" in meta_spec)
  2132       apply(drule_tac x="y" in meta_spec)
  2133 apply(auto)[1]
  2133       apply(auto)[1]
  2134 apply(rule_tac x="OrR1 <a>.M0 b" in exI)
  2134       apply(rule_tac x="OrR1 <a>.M0 b" in exI)
  2135 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2135       apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2136 apply(auto)[1]
  2136       apply(auto)[1]
  2137 (* OrR2 *)
  2137     (* OrR2 *)
  2138 apply(drule sym)
  2138      apply(drule sym)
  2139 apply(drule nrename_OrR2)
  2139      apply(drule nrename_OrR2)
  2140 apply(simp)
  2140       apply(simp)
  2141 apply(auto)[1]
  2141      apply(auto)[1]
  2142 apply(drule_tac x="N'" in meta_spec)
  2142      apply(drule_tac x="N'" in meta_spec)
  2143 apply(drule_tac x="x" in meta_spec)
  2143      apply(drule_tac x="x" in meta_spec)
  2144 apply(drule_tac x="y" in meta_spec)
  2144      apply(drule_tac x="y" in meta_spec)
  2145 apply(auto)[1]
  2145      apply(auto)[1]
  2146 apply(rule_tac x="OrR2 <a>.M0 b" in exI)
  2146      apply(rule_tac x="OrR2 <a>.M0 b" in exI)
  2147 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2147      apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2148 apply(auto)[1]
  2148      apply(auto)[1]
  2149 (* ImpL *)
  2149     (* ImpL *)
  2150 apply(drule sym)
  2150     apply(drule sym)
  2151 apply(frule nrename_ImpL_aux)
  2151     apply(frule nrename_ImpL_aux)
  2152 apply(erule disjE)
  2152     apply(erule disjE)
  2153 apply(auto)[1]
  2153      apply(auto)[1]
  2154 apply(drule nrename_ImpL')
  2154     apply(drule nrename_ImpL')
  2155 apply(simp add: fresh_prod fresh_atm)
  2155        apply(simp add: fresh_prod fresh_atm)
  2156 apply(simp add: fresh_atm)
  2156       apply(simp add: fresh_atm)
  2157 apply(simp add: fresh_atm)
  2157      apply(simp add: fresh_atm)
  2158 apply(erule disjE)
  2158     apply(erule disjE)
  2159 apply(auto)[1]
  2159      apply(auto)[1]
  2160 apply(drule_tac x="M'a" in meta_spec)
  2160      apply(drule_tac x="M'a" in meta_spec)
  2161 apply(drule_tac x="xa" in meta_spec)
  2161      apply(drule_tac x="xa" in meta_spec)
  2162 apply(drule_tac x="ya" in meta_spec)
  2162      apply(drule_tac x="ya" in meta_spec)
  2163 apply(auto)[1]
  2163      apply(auto)[1]
  2164 apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
  2164      apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
  2165 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2165      apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2166 apply(auto)[1]
  2166      apply(auto)[1]
  2167 apply(rule trans)
  2167      apply(rule trans)
  2168 apply(rule nrename.simps)
  2168       apply(rule nrename.simps)
  2169 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2169        apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2170 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2170       apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2171 apply(auto intro: fresh_a_redu)[1]
  2171       apply(auto intro: fresh_a_redu)[1]
  2172 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2172      apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2173 apply(auto)[1]
  2173     apply(auto)[1]
  2174 apply(drule_tac x="M'a" in meta_spec)
  2174     apply(drule_tac x="M'a" in meta_spec)
  2175 apply(drule_tac x="xa" in meta_spec)
  2175     apply(drule_tac x="xa" in meta_spec)
  2176 apply(drule_tac x="y" in meta_spec)
  2176     apply(drule_tac x="y" in meta_spec)
  2177 apply(auto)[1]
  2177     apply(auto)[1]
  2178 apply(rule_tac x="ImpL <a>.M0 (x).N' xa" in exI)
  2178     apply(rule_tac x="ImpL <a>.M0 (x).N' xa" in exI)
  2179 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2179     apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2180 apply(auto)[1]
  2180     apply(auto)[1]
  2181 apply(rule trans)
  2181     apply(rule trans)
  2182 apply(rule nrename.simps)
  2182      apply(rule nrename.simps)
  2183 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2183       apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2184 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2184      apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2185 apply(auto intro: fresh_a_redu)[1]
  2185      apply(auto intro: fresh_a_redu)[1]
  2186 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2186     apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2187 apply(drule sym)
  2187    apply(drule sym)
  2188 apply(frule nrename_ImpL_aux)
  2188    apply(frule nrename_ImpL_aux)
  2189 apply(erule disjE)
  2189    apply(erule disjE)
  2190 apply(auto)[1]
  2190     apply(auto)[1]
  2191 apply(drule nrename_ImpL')
  2191    apply(drule nrename_ImpL')
  2192 apply(simp add: fresh_prod fresh_atm)
  2192       apply(simp add: fresh_prod fresh_atm)
  2193 apply(simp add: fresh_atm)
  2193      apply(simp add: fresh_atm)
  2194 apply(simp add: fresh_atm)
  2194     apply(simp add: fresh_atm)
  2195 apply(erule disjE)
  2195    apply(erule disjE)
  2196 apply(auto)[1]
  2196     apply(auto)[1]
  2197 apply(drule_tac x="N'a" in meta_spec)
  2197     apply(drule_tac x="N'a" in meta_spec)
  2198 apply(drule_tac x="xa" in meta_spec)
  2198     apply(drule_tac x="xa" in meta_spec)
  2199 apply(drule_tac x="ya" in meta_spec)
  2199     apply(drule_tac x="ya" in meta_spec)
  2200 apply(auto)[1]
  2200     apply(auto)[1]
  2201 apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
  2201     apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
  2202 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2202     apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2203 apply(auto)[1]
  2203     apply(auto)[1]
  2204 apply(rule trans)
  2204     apply(rule trans)
  2205 apply(rule nrename.simps)
  2205      apply(rule nrename.simps)
  2206 apply(auto intro: fresh_a_redu)[1]
  2206       apply(auto intro: fresh_a_redu)[1]
  2207 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2207      apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2208 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2208     apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2209 apply(auto)[1]
  2209    apply(auto)[1]
  2210 apply(drule_tac x="N'a" in meta_spec)
  2210    apply(drule_tac x="N'a" in meta_spec)
  2211 apply(drule_tac x="xa" in meta_spec)
  2211    apply(drule_tac x="xa" in meta_spec)
  2212 apply(drule_tac x="y" in meta_spec)
  2212    apply(drule_tac x="y" in meta_spec)
  2213 apply(auto)[1]
  2213    apply(auto)[1]
  2214 apply(rule_tac x="ImpL <a>.M' (x).M0 xa" in exI)
  2214    apply(rule_tac x="ImpL <a>.M' (x).M0 xa" in exI)
  2215 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2215    apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2216 apply(auto)[1]
  2216    apply(auto)[1]
  2217 apply(rule trans)
  2217    apply(rule trans)
  2218 apply(rule nrename.simps)
  2218     apply(rule nrename.simps)
  2219 apply(auto intro: fresh_a_redu)[1]
  2219      apply(auto intro: fresh_a_redu)[1]
  2220 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2220     apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2221 apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2221    apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
  2222 (* ImpR *)
  2222     (* ImpR *)
  2223 apply(drule sym)
  2223   apply(drule sym)
  2224 apply(drule nrename_ImpR)
  2224   apply(drule nrename_ImpR)
  2225 apply(simp)
  2225     apply(simp)
  2226 apply(simp)
  2226    apply(simp)
  2227 apply(auto)[1]
  2227   apply(auto)[1]
  2228 apply(drule_tac x="N'" in meta_spec)
  2228   apply(drule_tac x="N'" in meta_spec)
  2229 apply(drule_tac x="xa" in meta_spec)
  2229   apply(drule_tac x="xa" in meta_spec)
  2230 apply(drule_tac x="y" in meta_spec)
  2230   apply(drule_tac x="y" in meta_spec)
  2231 apply(auto)[1]
  2231   apply(auto)[1]
  2232 apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
  2232   apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
  2233 apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2233   apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
  2234 apply(auto)[1]
  2234   apply(auto)[1]
  2235 done
  2235   done
  2236 
  2236 
  2237 lemma SNa_preserved_renaming2:
  2237 lemma SNa_preserved_renaming2:
  2238   assumes a: "SNa N"
  2238   assumes a: "SNa N"
  2239   shows "SNa (N[x\<turnstile>n>y])"
  2239   shows "SNa (N[x\<turnstile>n>y])"
  2240 using a
  2240   using a
  2241 apply(induct rule: SNa_induct)
  2241   apply(induct rule: SNa_induct)
  2242 apply(case_tac "x=y")
  2242   apply(case_tac "x=y")
  2243 apply(simp add: nrename_id)
  2243    apply(simp add: nrename_id)
  2244 apply(rule SNaI)
  2244   apply(rule SNaI)
  2245 apply(drule nrename_aredu)
  2245   apply(drule nrename_aredu)
  2246 apply(blast)+
  2246    apply(blast)+
  2247 done
  2247   done
  2248 
  2248 
  2249 text \<open>helper-stuff to set up the induction\<close>
  2249 text \<open>helper-stuff to set up the induction\<close>
  2250 
  2250 
  2251 abbreviation
  2251 abbreviation
  2252   SNa_set :: "trm set"
  2252   SNa_set :: "trm set"
  2253 where
  2253   where
  2254   "SNa_set \<equiv> {M. SNa M}"
  2254     "SNa_set \<equiv> {M. SNa M}"
  2255 
  2255 
  2256 abbreviation
  2256 abbreviation
  2257   A_Redu_set :: "(trm\<times>trm) set"
  2257   A_Redu_set :: "(trm\<times>trm) set"
  2258 where
  2258   where
  2259  "A_Redu_set \<equiv> {(N,M)| M N. M \<longrightarrow>\<^sub>a N}"
  2259     "A_Redu_set \<equiv> {(N,M)| M N. M \<longrightarrow>\<^sub>a N}"
  2260 
  2260 
  2261 lemma SNa_elim:
  2261 lemma SNa_elim:
  2262   assumes a: "SNa M"
  2262   assumes a: "SNa M"
  2263   shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^sub>a N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
  2263   shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^sub>a N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
  2264 using a
  2264   using a
  2265 by (induct rule: SNa.induct) (blast)
  2265   by (induct rule: SNa.induct) (blast)
  2266 
  2266 
  2267 lemma wf_SNa_restricted:
  2267 lemma wf_SNa_restricted:
  2268   shows "wf (A_Redu_set \<inter> (UNIV \<times> SNa_set))"
  2268   shows "wf (A_Redu_set \<inter> (UNIV \<times> SNa_set))"
  2269 apply(unfold wf_def)
  2269   apply(unfold wf_def)
  2270 apply(intro strip)
  2270   apply(intro strip)
  2271 apply(case_tac "SNa x")
  2271   apply(case_tac "SNa x")
  2272 apply(simp (no_asm_use))
  2272    apply(simp (no_asm_use))
  2273 apply(drule_tac P="P" in SNa_elim)
  2273    apply(drule_tac P="P" in SNa_elim)
  2274 apply(erule mp)
  2274    apply(erule mp)
  2275 apply(blast)
  2275    apply(blast)
  2276 (* other case *)
  2276     (* other case *)
  2277 apply(drule_tac x="x" in spec)
  2277   apply(drule_tac x="x" in spec)
  2278 apply(erule mp)
  2278   apply(erule mp)
  2279 apply(fast)
  2279   apply(fast)
  2280 done
  2280   done
  2281 
  2281 
  2282 definition SNa_Redu :: "(trm \<times> trm) set" where
  2282 definition SNa_Redu :: "(trm \<times> trm) set" where
  2283   "SNa_Redu \<equiv> A_Redu_set \<inter> (UNIV \<times> SNa_set)"
  2283   "SNa_Redu \<equiv> A_Redu_set \<inter> (UNIV \<times> SNa_set)"
  2284 
  2284 
  2285 lemma wf_SNa_Redu:
  2285 lemma wf_SNa_Redu:
  2286   shows "wf SNa_Redu"
  2286   shows "wf SNa_Redu"
  2287 apply(unfold SNa_Redu_def)
  2287   apply(unfold SNa_Redu_def)
  2288 apply(rule wf_SNa_restricted)
  2288   apply(rule wf_SNa_restricted)
  2289 done
  2289   done
  2290 
  2290 
  2291 lemma wf_measure_triple:
  2291 lemma wf_measure_triple:
  2292 shows "wf ((measure size) <*lex*> SNa_Redu <*lex*> SNa_Redu)"
  2292   shows "wf ((measure size) <*lex*> SNa_Redu <*lex*> SNa_Redu)"
  2293 by (auto intro: wf_SNa_Redu)
  2293   by (auto intro: wf_SNa_Redu)
  2294 
  2294 
  2295 lemma my_wf_induct_triple: 
  2295 lemma my_wf_induct_triple: 
  2296  assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
  2296   assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
  2297  and     b: "\<And>x. \<lbrakk>\<And>y. ((fst y,fst (snd y),snd (snd y)),(fst x,fst (snd x), snd (snd x))) 
  2297     and     b: "\<And>x. \<lbrakk>\<And>y. ((fst y,fst (snd y),snd (snd y)),(fst x,fst (snd x), snd (snd x))) 
  2298                                     \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y\<rbrakk> \<Longrightarrow> P x"  
  2298                                     \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y\<rbrakk> \<Longrightarrow> P x"  
  2299  shows "P x"
  2299   shows "P x"
  2300 using a
  2300   using a
  2301 apply(induct x rule: wf_induct_rule)
  2301   apply(induct x rule: wf_induct_rule)
  2302 apply(rule b)
  2302   apply(rule b)
  2303 apply(simp)
  2303   apply(simp)
  2304 done
  2304   done
  2305 
  2305 
  2306 lemma my_wf_induct_triple': 
  2306 lemma my_wf_induct_triple': 
  2307  assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
  2307   assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
  2308  and    b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P (y1,y2,y3)\<rbrakk> 
  2308     and    b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P (y1,y2,y3)\<rbrakk> 
  2309              \<Longrightarrow> P (x1,x2,x3)"  
  2309              \<Longrightarrow> P (x1,x2,x3)"  
  2310  shows "P (x1,x2,x3)"
  2310   shows "P (x1,x2,x3)"
  2311 apply(rule_tac my_wf_induct_triple[OF a])
  2311   apply(rule_tac my_wf_induct_triple[OF a])
  2312 apply(case_tac x rule: prod.exhaust)
  2312   apply(case_tac x rule: prod.exhaust)
  2313 apply(simp)
  2313   apply(simp)
  2314 apply(rename_tac p a b)
  2314   apply(rename_tac p a b)
  2315 apply(case_tac b)
  2315   apply(case_tac b)
  2316 apply(simp)
  2316   apply(simp)
  2317 apply(rule b)
  2317   apply(rule b)
  2318 apply(blast)
  2318   apply(blast)
  2319 done
  2319   done
  2320 
  2320 
  2321 lemma my_wf_induct_triple'': 
  2321 lemma my_wf_induct_triple'': 
  2322  assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
  2322   assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
  2323  and     b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y1 y2 y3\<rbrakk>
  2323     and     b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y1 y2 y3\<rbrakk>
  2324                \<Longrightarrow> P x1 x2 x3"  
  2324                \<Longrightarrow> P x1 x2 x3"  
  2325  shows "P x1 x2 x3"
  2325   shows "P x1 x2 x3"
  2326 apply(rule_tac my_wf_induct_triple'[where P="\<lambda>(x1,x2,x3). P x1 x2 x3", simplified])
  2326   apply(rule_tac my_wf_induct_triple'[where P="\<lambda>(x1,x2,x3). P x1 x2 x3", simplified])
  2327 apply(rule a)
  2327    apply(rule a)
  2328 apply(rule b)
  2328   apply(rule b)
  2329 apply(auto)
  2329   apply(auto)
  2330 done
  2330   done
  2331 
  2331 
  2332 lemma excluded_m:
  2332 lemma excluded_m:
  2333   assumes a: "<a>:M \<in> (\<parallel><B>\<parallel>)" "(x):N \<in> (\<parallel>(B)\<parallel>)"
  2333   assumes a: "<a>:M \<in> (\<parallel><B>\<parallel>)" "(x):N \<in> (\<parallel>(B)\<parallel>)"
  2334   shows "(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))
  2334   shows "(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))
  2335       \<or>\<not>(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))"
  2335       \<or>\<not>(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))"
  2336 by (blast)
  2336   by (blast)
       
  2337 
       
  2338 text \<open>The following two simplification rules are necessary because of the 
       
  2339       new definition of lexicographic ordering\<close>
       
  2340 lemma ne_and_SNa_Redu[simp]: "M \<noteq> x \<and> (M,x) \<in> SNa_Redu \<longleftrightarrow> (M,x) \<in> SNa_Redu"
       
  2341   using wf_SNa_Redu by auto
       
  2342 
       
  2343 lemma ne_and_less_size [simp]: "A \<noteq> B \<and> size A < size B \<longleftrightarrow> size A < size B"
       
  2344   by auto
  2337 
  2345 
  2338 lemma tricky_subst:
  2346 lemma tricky_subst:
  2339   assumes a1: "b\<sharp>(c,N)"
  2347   assumes a1: "b\<sharp>(c,N)"
  2340   and     a2: "z\<sharp>(x,P)"
  2348     and     a2: "z\<sharp>(x,P)"
  2341   and     a3: "M\<noteq>Ax z b"
  2349     and     a3: "M\<noteq>Ax z b"
  2342   shows "(Cut <c>.N (z).M){b:=(x).P} = Cut <c>.N (z).(M{b:=(x).P})"
  2350   shows "(Cut <c>.N (z).M){b:=(x).P} = Cut <c>.N (z).(M{b:=(x).P})"
  2343 using a1 a2 a3
  2351   using a1 a2 a3
  2344 apply -
  2352   apply -
  2345 apply(generate_fresh "coname")
  2353   apply(generate_fresh "coname")
  2346 apply(subgoal_tac "Cut <c>.N (z).M = Cut <ca>.([(ca,c)]\<bullet>N) (z).M")
  2354   apply(subgoal_tac "Cut <c>.N (z).M = Cut <ca>.([(ca,c)]\<bullet>N) (z).M")
  2347 apply(simp)
  2355    apply(simp)
  2348 apply(rule trans)
  2356    apply(rule trans)
  2349 apply(rule better_Cut_substc)
  2357     apply(rule better_Cut_substc)
  2350 apply(simp)
  2358      apply(simp)
  2351 apply(simp add: abs_fresh)
  2359     apply(simp add: abs_fresh)
  2352 apply(simp)
  2360    apply(simp)
  2353 apply(subgoal_tac "b\<sharp>([(ca,c)]\<bullet>N)")
  2361    apply(subgoal_tac "b\<sharp>([(ca,c)]\<bullet>N)")
  2354 apply(simp add: forget)
  2362     apply(simp add: forget)
  2355 apply(simp add: trm.inject)
  2363     apply(simp add: trm.inject)
  2356 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
  2364    apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
  2357 apply(simp add: trm.inject)
  2365   apply(simp add: trm.inject)
  2358 apply(rule sym)
  2366   apply(rule sym)
  2359 apply(simp add: alpha fresh_prod fresh_atm)
  2367   apply(simp add: alpha fresh_prod fresh_atm)
  2360 done
  2368   done
  2361 
  2369 
  2362 text \<open>3rd lemma\<close>
  2370 text \<open>3rd lemma\<close>
  2363 
  2371 
  2364 lemma CUT_SNa_aux:
  2372 lemma CUT_SNa_aux:
  2365   assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
  2373   assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
  2366   and     a2: "SNa M"
  2374     and     a2: "SNa M"
  2367   and     a3: "(x):N \<in> (\<parallel>(B)\<parallel>)"
  2375     and     a3: "(x):N \<in> (\<parallel>(B)\<parallel>)"
  2368   and     a4: "SNa N"
  2376     and     a4: "SNa N"
  2369   shows   "SNa (Cut <a>.M (x).N)"
  2377   shows   "SNa (Cut <a>.M (x).N)"
  2370 using a1 a2 a3 a4 [[simproc del: defined_all]]
  2378   using a1 a2 a3 a4 [[simproc del: defined_all]]
  2371 apply(induct B M N arbitrary: a x rule: my_wf_induct_triple''[OF wf_measure_triple])
  2379   apply(induct B M N arbitrary: a x rule: my_wf_induct_triple''[OF wf_measure_triple])
  2372 apply(rule SNaI)
  2380   apply(rule SNaI)
  2373 apply(drule Cut_a_redu_elim)
  2381   apply(drule Cut_a_redu_elim)
  2374 apply(erule disjE)
  2382   apply(erule disjE)
  2375 (* left-inner reduction *)
  2383     (* left-inner reduction *)
  2376 apply(erule exE)
  2384    apply(erule exE)
  2377 apply(erule conjE)+
  2385    apply(erule conjE)+
  2378 apply(simp)
  2386    apply(simp)
  2379 apply(drule_tac x="x1" in meta_spec)
  2387    apply(drule_tac x="x1" in meta_spec)
  2380 apply(drule_tac x="M'a" in meta_spec)
  2388    apply(drule_tac x="M'a" in meta_spec)
  2381 apply(drule_tac x="x3" in meta_spec)
  2389    apply(drule_tac x="x3" in meta_spec)
  2382 apply(drule conjunct2)
  2390    apply(drule conjunct2)
  2383 apply(drule mp)
  2391    apply(drule mp)
  2384 apply(rule conjI)
  2392     apply(rule conjI)
  2385 apply(simp)
  2393      apply(simp)
  2386 apply(rule disjI1)
  2394     apply(rule disjI1)
  2387 apply(simp add: SNa_Redu_def)
  2395     apply(simp add: SNa_Redu_def)
  2388 apply(drule_tac x="a" in spec)
  2396    apply(drule_tac x="a" in spec)
  2389 apply(drule mp)
  2397    apply(drule mp)
  2390 apply(simp add: CANDs_preserved_single)
  2398     apply(simp add: CANDs_preserved_single)
  2391 apply(drule mp)
  2399    apply(drule mp)
  2392 apply(simp add: a_preserves_SNa)
  2400     apply(simp add: a_preserves_SNa)
  2393 apply(drule_tac x="x" in spec)
  2401    apply(drule_tac x="x" in spec)
  2394 apply(simp)
  2402    apply(simp)
  2395 apply(erule disjE)
  2403   apply(erule disjE)
  2396 (* right-inner reduction *)
  2404     (* right-inner reduction *)
  2397 apply(erule exE)
  2405    apply(erule exE)
  2398 apply(erule conjE)+
  2406    apply(erule conjE)+
  2399 apply(simp)
  2407    apply(simp)
  2400 apply(drule_tac x="x1" in meta_spec)
  2408    apply(drule_tac x="x1" in meta_spec)
  2401 apply(drule_tac x="x2" in meta_spec)
  2409    apply(drule_tac x="x2" in meta_spec)
  2402 apply(drule_tac x="N'" in meta_spec)
  2410    apply(drule_tac x="N'" in meta_spec)
  2403 apply(drule conjunct2)
  2411    apply(drule conjunct2)
  2404 apply(drule mp)
  2412    apply(drule mp)
  2405 apply(rule conjI)
  2413     apply(rule conjI)
  2406 apply(simp)
  2414      apply(simp)
  2407 apply(rule disjI2)
  2415     apply(rule disjI2)
  2408 apply(simp add: SNa_Redu_def)
  2416     apply(simp add: SNa_Redu_def)
  2409 apply(drule_tac x="a" in spec)
  2417    apply(drule_tac x="a" in spec)
  2410 apply(drule mp)
  2418    apply(drule mp)
  2411 apply(simp add: CANDs_preserved_single)
  2419     apply(simp add: CANDs_preserved_single)
  2412 apply(drule mp)
  2420    apply(drule mp)
  2413 apply(assumption)
  2421     apply(assumption)
  2414 apply(drule_tac x="x" in spec)
  2422    apply(drule_tac x="x" in spec)
  2415 apply(drule mp)
  2423    apply(drule mp)
  2416 apply(simp add: CANDs_preserved_single)
  2424     apply(simp add: CANDs_preserved_single)
  2417 apply(drule mp)
  2425    apply(drule mp)
  2418 apply(simp add: a_preserves_SNa)
  2426     apply(simp add: a_preserves_SNa)
  2419 apply(assumption)
  2427    apply(assumption)
  2420 apply(erule disjE)
  2428   apply(erule disjE)
  2421 (******** c-reduction *********)
  2429     (******** c-reduction *********)
  2422 apply(drule Cut_c_redu_elim)
  2430    apply(drule Cut_c_redu_elim)
  2423 (* c-left reduction*)
  2431     (* c-left reduction*)
  2424 apply(erule disjE)
  2432    apply(erule disjE)
  2425 apply(erule conjE)
  2433     apply(erule conjE)
  2426 apply(frule_tac B="x1" in fic_CANDS)
  2434     apply(frule_tac B="x1" in fic_CANDS)
  2427 apply(simp)
  2435      apply(simp)
  2428 apply(erule disjE)
  2436     apply(erule disjE)
  2429 (* in AXIOMSc *)
  2437     (* in AXIOMSc *)
  2430 apply(simp add: AXIOMSc_def)
  2438      apply(simp add: AXIOMSc_def)
  2431 apply(erule exE)+
  2439      apply(erule exE)+
  2432 apply(simp add: ctrm.inject)
  2440      apply(simp add: ctrm.inject)
  2433 apply(simp add: alpha)
  2441      apply(simp add: alpha)
  2434 apply(erule disjE)
  2442      apply(erule disjE)
  2435 apply(simp)
  2443       apply(simp)
  2436 apply(rule impI)
  2444       apply(rule impI)
  2437 apply(simp)
  2445       apply(simp)
  2438 apply(subgoal_tac "fic (Ax y b) b")(*A*)
  2446       apply(subgoal_tac "fic (Ax y b) b")(*A*)
  2439 apply(simp)
  2447        apply(simp)
  2440 (*A*)
  2448     (*A*)
  2441 apply(auto)[1]
  2449       apply(auto)[1]
  2442 apply(simp)
  2450      apply(simp)
  2443 apply(rule impI)
  2451      apply(rule impI)
  2444 apply(simp)
  2452      apply(simp)
  2445 apply(subgoal_tac "fic (Ax ([(a,aa)]\<bullet>y) a) a")(*B*)
  2453      apply(subgoal_tac "fic (Ax ([(a,aa)]\<bullet>y) a) a")(*B*)
  2446 apply(simp)
  2454       apply(simp)
  2447 (*B*)
  2455     (*B*)
  2448 apply(auto)[1]
  2456      apply(auto)[1]
  2449 (* in BINDINGc *)
  2457     (* in BINDINGc *)
  2450 apply(simp)
  2458     apply(simp)
  2451 apply(drule BINDINGc_elim)
  2459     apply(drule BINDINGc_elim)
  2452 apply(simp)
  2460     apply(simp)
  2453 (* c-right reduction*)
  2461     (* c-right reduction*)
  2454 apply(erule conjE)
  2462    apply(erule conjE)
  2455 apply(frule_tac B="x1" in fin_CANDS)
  2463    apply(frule_tac B="x1" in fin_CANDS)
  2456 apply(simp)
  2464     apply(simp)
  2457 apply(erule disjE)
  2465    apply(erule disjE)
  2458 (* in AXIOMSc *)
  2466     (* in AXIOMSc *)
  2459 apply(simp add: AXIOMSn_def)
  2467     apply(simp add: AXIOMSn_def)
  2460 apply(erule exE)+
  2468     apply(erule exE)+
  2461 apply(simp add: ntrm.inject)
  2469     apply(simp add: ntrm.inject)
  2462 apply(simp add: alpha)
  2470     apply(simp add: alpha)
  2463 apply(erule disjE)
  2471     apply(erule disjE)
  2464 apply(simp)
  2472      apply(simp)
  2465 apply(rule impI)
  2473      apply(rule impI)
  2466 apply(simp)
  2474      apply(simp)
  2467 apply(subgoal_tac "fin (Ax xa b) xa")(*A*)
  2475      apply(subgoal_tac "fin (Ax xa b) xa")(*A*)
  2468 apply(simp)
  2476       apply(simp)
  2469 (*A*)
  2477     (*A*)
  2470 apply(auto)[1]
  2478      apply(auto)[1]
  2471 apply(simp)
  2479     apply(simp)
  2472 apply(rule impI)
  2480     apply(rule impI)
  2473 apply(simp)
  2481     apply(simp)
  2474 apply(subgoal_tac "fin (Ax x ([(x,xa)]\<bullet>b)) x")(*B*)
  2482     apply(subgoal_tac "fin (Ax x ([(x,xa)]\<bullet>b)) x")(*B*)
  2475 apply(simp)
  2483      apply(simp)
  2476 (*B*)
  2484     (*B*)
  2477 apply(auto)[1]
  2485     apply(auto)[1]
  2478 (* in BINDINGc *)
  2486     (* in BINDINGc *)
  2479 apply(simp)
  2487    apply(simp)
  2480 apply(drule BINDINGn_elim)
  2488    apply(drule BINDINGn_elim)
  2481 apply(simp)
  2489    apply(simp)
  2482 (*********** l-reductions ************)
  2490     (*********** l-reductions ************)
  2483 apply(drule Cut_l_redu_elim)
  2491   apply(drule Cut_l_redu_elim)
  2484 apply(erule disjE)
  2492   apply(erule disjE)
  2485 (* ax1 *)
  2493     (* ax1 *)
  2486 apply(erule exE)
  2494    apply(erule exE)
  2487 apply(simp)
  2495    apply(simp)
  2488 apply(simp add: SNa_preserved_renaming1)
  2496    apply(simp add: SNa_preserved_renaming1)
  2489 apply(erule disjE)
  2497   apply(erule disjE)
  2490 (* ax2 *)
  2498     (* ax2 *)
  2491 apply(erule exE)
  2499    apply(erule exE)
  2492 apply(simp add: SNa_preserved_renaming2)
  2500    apply(simp add: SNa_preserved_renaming2)
  2493 apply(erule disjE)
  2501   apply(erule disjE)
  2494 (* LNot *)
  2502     (* LNot *)
  2495 apply(erule exE)+
  2503    apply(erule exE)+
  2496 apply(auto)[1]
  2504    apply(auto)[1]
  2497 apply(frule_tac excluded_m)
  2505    apply(frule_tac excluded_m)
  2498 apply(assumption)
  2506     apply(assumption)
  2499 apply(erule disjE)
  2507    apply(erule disjE)
  2500 (* one of them in BINDING *)
  2508     (* one of them in BINDING *)
  2501 apply(erule disjE)
  2509     apply(erule disjE)
  2502 apply(drule fin_elims)
  2510      apply(drule fin_elims)
  2503 apply(drule fic_elims)
  2511      apply(drule fic_elims)
  2504 apply(simp)
  2512      apply(simp)
  2505 apply(drule BINDINGc_elim)
  2513      apply(drule BINDINGc_elim)
  2506 apply(drule_tac x="x" in spec)
  2514      apply(drule_tac x="x" in spec)
  2507 apply(drule_tac x="NotL <b>.N' x" in spec)
  2515      apply(drule_tac x="NotL <b>.N' x" in spec)
  2508 apply(simp)
  2516      apply(simp)
  2509 apply(simp add: better_NotR_substc)
  2517      apply(simp add: better_NotR_substc)
  2510 apply(generate_fresh "coname")
  2518      apply(generate_fresh "coname")
  2511 apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.NotR (y).M'a a' (x).NotL <b>.N' x) 
  2519      apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.NotR (y).M'a a' (x).NotL <b>.N' x) 
  2512                    =  Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x")
  2520                    =  Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x")
  2513 apply(simp)
  2521       apply(simp)
  2514 apply(subgoal_tac "Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
  2522       apply(subgoal_tac "Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
  2515 apply(simp only: a_preserves_SNa)
  2523        apply(simp only: a_preserves_SNa)
  2516 apply(rule al_redu)
  2524       apply(rule al_redu)
  2517 apply(rule better_LNot_intro)
  2525       apply(rule better_LNot_intro)
  2518 apply(simp)
  2526        apply(simp)
  2519 apply(simp)
  2527       apply(simp)
  2520 apply(fresh_fun_simp (no_asm))
  2528      apply(fresh_fun_simp (no_asm))
  2521 apply(simp)
  2529      apply(simp)
  2522 (* other case of in BINDING *)
  2530     (* other case of in BINDING *)
  2523 apply(drule fin_elims)
  2531     apply(drule fin_elims)
  2524 apply(drule fic_elims)
  2532     apply(drule fic_elims)
  2525 apply(simp)
  2533     apply(simp)
  2526 apply(drule BINDINGn_elim)
  2534     apply(drule BINDINGn_elim)
  2527 apply(drule_tac x="a" in spec)
  2535     apply(drule_tac x="a" in spec)
  2528 apply(drule_tac x="NotR (y).M'a a" in spec)
  2536     apply(drule_tac x="NotR (y).M'a a" in spec)
  2529 apply(simp)
  2537     apply(simp)
  2530 apply(simp add: better_NotL_substn)
  2538     apply(simp add: better_NotL_substn)
  2531 apply(generate_fresh "name")
  2539     apply(generate_fresh "name")
  2532 apply(subgoal_tac "fresh_fun (\<lambda>x'. Cut <a>.NotR (y).M'a a (x').NotL <b>.N' x') 
  2540     apply(subgoal_tac "fresh_fun (\<lambda>x'. Cut <a>.NotR (y).M'a a (x').NotL <b>.N' x') 
  2533                    = Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c")
  2541                    = Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c")
  2534 apply(simp)
  2542      apply(simp)
  2535 apply(subgoal_tac "Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
  2543      apply(subgoal_tac "Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
  2536 apply(simp only: a_preserves_SNa)
  2544       apply(simp only: a_preserves_SNa)
  2537 apply(rule al_redu)
  2545      apply(rule al_redu)
  2538 apply(rule better_LNot_intro)
  2546      apply(rule better_LNot_intro)
  2539 apply(simp)
  2547       apply(simp)
  2540 apply(simp)
  2548      apply(simp)
  2541 apply(fresh_fun_simp (no_asm))
  2549     apply(fresh_fun_simp (no_asm))
  2542 apply(simp)
  2550     apply(simp)
  2543 (* none of them in BINDING *)
  2551     (* none of them in BINDING *)
  2544 apply(simp)
  2552    apply(simp)
  2545 apply(erule conjE)
  2553    apply(erule conjE)
  2546 apply(frule CAND_NotR_elim)
  2554    apply(frule CAND_NotR_elim)
  2547 apply(assumption)
  2555     apply(assumption)
  2548 apply(erule exE)
  2556    apply(erule exE)
  2549 apply(frule CAND_NotL_elim)
  2557    apply(frule CAND_NotL_elim)
  2550 apply(assumption)
  2558     apply(assumption)
  2551 apply(erule exE)
  2559    apply(erule exE)
  2552 apply(simp only: ty.inject)
  2560    apply(simp only: ty.inject)
  2553 apply(drule_tac x="B'" in meta_spec)
  2561    apply(drule_tac x="B'" in meta_spec)
  2554 apply(drule_tac x="N'" in meta_spec)
  2562    apply(drule_tac x="N'" in meta_spec)
  2555 apply(drule_tac x="M'a" in meta_spec)
  2563    apply(drule_tac x="M'a" in meta_spec)
  2556 apply(erule conjE)+
  2564    apply(erule conjE)+
  2557 apply(drule mp)
  2565    apply(drule mp)
  2558 apply(simp)
  2566     apply(simp)
  2559 apply(drule_tac x="b" in spec)
  2567    apply(drule_tac x="b" in spec)
  2560 apply(rotate_tac 13)
  2568    apply(rotate_tac 13)
  2561 apply(drule mp)
  2569    apply(drule mp)
  2562 apply(assumption)
  2570     apply(assumption)
  2563 apply(rotate_tac 13)
  2571    apply(rotate_tac 13)
  2564 apply(drule mp)
  2572    apply(drule mp)
  2565 apply(simp add: CANDs_imply_SNa)
  2573     apply(simp add: CANDs_imply_SNa)
  2566 apply(drule_tac x="y" in spec)
  2574    apply(drule_tac x="y" in spec)
  2567 apply(rotate_tac 13)
  2575    apply(rotate_tac 13)
  2568 apply(drule mp)
  2576    apply(drule mp)
  2569 apply(assumption)
  2577     apply(assumption)
  2570 apply(rotate_tac 13)
  2578    apply(rotate_tac 13)
  2571 apply(drule mp)
  2579    apply(drule mp)
  2572 apply(simp add: CANDs_imply_SNa)
  2580     apply(simp add: CANDs_imply_SNa)
  2573 apply(assumption)
  2581    apply(assumption)
  2574 (* LAnd1 case *)
  2582     (* LAnd1 case *)
  2575 apply(erule disjE)
  2583   apply(erule disjE)
  2576 apply(erule exE)+
  2584    apply(erule exE)+
  2577 apply(auto)[1]
  2585    apply(auto)[1]
  2578 apply(frule_tac excluded_m)
  2586    apply(frule_tac excluded_m)
  2579 apply(assumption)
  2587     apply(assumption)
  2580 apply(erule disjE)
  2588    apply(erule disjE)
  2581 (* one of them in BINDING *)
  2589     (* one of them in BINDING *)
  2582 apply(erule disjE)
  2590     apply(erule disjE)
  2583 apply(drule fin_elims)
  2591      apply(drule fin_elims)
  2584 apply(drule fic_elims)
  2592      apply(drule fic_elims)
  2585 apply(simp)
  2593      apply(simp)
  2586 apply(drule BINDINGc_elim)
  2594      apply(drule BINDINGc_elim)
  2587 apply(drule_tac x="x" in spec)
  2595      apply(drule_tac x="x" in spec)
  2588 apply(drule_tac x="AndL1 (y).N' x" in spec)
  2596      apply(drule_tac x="AndL1 (y).N' x" in spec)
  2589 apply(simp)
  2597      apply(simp)
  2590 apply(simp add: better_AndR_substc)
  2598      apply(simp add: better_AndR_substc)
  2591 apply(generate_fresh "coname")
  2599      apply(generate_fresh "coname")
  2592 apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL1 (y).N' x) 
  2600      apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL1 (y).N' x) 
  2593                    = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x")
  2601                    = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x")
  2594 apply(simp)
  2602       apply(simp)
  2595 apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
  2603       apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
  2596 apply(auto intro: a_preserves_SNa)[1]
  2604        apply(auto intro: a_preserves_SNa)[1]
  2597 apply(rule al_redu)
  2605       apply(rule al_redu)
  2598 apply(rule better_LAnd1_intro)
  2606       apply(rule better_LAnd1_intro)
  2599 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2607        apply(simp add: abs_fresh fresh_prod fresh_atm)
  2600 apply(simp)
  2608       apply(simp)
  2601 apply(fresh_fun_simp (no_asm))
  2609      apply(fresh_fun_simp (no_asm))
  2602 apply(simp)
  2610      apply(simp)
  2603 (* other case of in BINDING *)
  2611     (* other case of in BINDING *)
  2604 apply(drule fin_elims)
  2612     apply(drule fin_elims)
  2605 apply(drule fic_elims)
  2613     apply(drule fic_elims)
  2606 apply(simp)
  2614     apply(simp)
  2607 apply(drule BINDINGn_elim)
  2615     apply(drule BINDINGn_elim)
  2608 apply(drule_tac x="a" in spec)
  2616     apply(drule_tac x="a" in spec)
  2609 apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
  2617     apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
  2610 apply(simp)
  2618     apply(simp)
  2611 apply(simp add: better_AndL1_substn)
  2619     apply(simp add: better_AndL1_substn)
  2612 apply(generate_fresh "name")
  2620     apply(generate_fresh "name")
  2613 apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL1 (y).N' z') 
  2621     apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL1 (y).N' z') 
  2614                    = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca")
  2622                    = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca")
  2615 apply(simp)
  2623      apply(simp)
  2616 apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
  2624      apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
  2617 apply(auto intro: a_preserves_SNa)[1]
  2625       apply(auto intro: a_preserves_SNa)[1]
  2618 apply(rule al_redu)
  2626      apply(rule al_redu)
  2619 apply(rule better_LAnd1_intro)
  2627      apply(rule better_LAnd1_intro)
  2620 apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2628       apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2621 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2629      apply(simp add: abs_fresh fresh_prod fresh_atm)
  2622 apply(fresh_fun_simp (no_asm))
  2630     apply(fresh_fun_simp (no_asm))
  2623 apply(simp)
  2631     apply(simp)
  2624 (* none of them in BINDING *)
  2632     (* none of them in BINDING *)
  2625 apply(simp)
  2633    apply(simp)
  2626 apply(erule conjE)
  2634    apply(erule conjE)
  2627 apply(frule CAND_AndR_elim)
  2635    apply(frule CAND_AndR_elim)
  2628 apply(assumption)
  2636     apply(assumption)
  2629 apply(erule exE)
  2637    apply(erule exE)
  2630 apply(frule CAND_AndL1_elim)
  2638    apply(frule CAND_AndL1_elim)
  2631 apply(assumption)
  2639     apply(assumption)
  2632 apply(erule exE)+
  2640    apply(erule exE)+
  2633 apply(simp only: ty.inject)
  2641    apply(simp only: ty.inject)
  2634 apply(drule_tac x="B1" in meta_spec)
  2642    apply(drule_tac x="B1" in meta_spec)
  2635 apply(drule_tac x="M1" in meta_spec)
  2643    apply(drule_tac x="M1" in meta_spec)
  2636 apply(drule_tac x="N'" in meta_spec)
  2644    apply(drule_tac x="N'" in meta_spec)
  2637 apply(erule conjE)+
  2645    apply(erule conjE)+
  2638 apply(drule mp)
  2646    apply(drule mp)
  2639 apply(simp)
  2647     apply(simp)
  2640 apply(drule_tac x="b" in spec)
  2648    apply(drule_tac x="b" in spec)
  2641 apply(rotate_tac 14)
  2649    apply(rotate_tac 14)
  2642 apply(drule mp)
  2650    apply(drule mp)
  2643 apply(assumption)
  2651     apply(assumption)
  2644 apply(rotate_tac 14)
  2652    apply(rotate_tac 14)
  2645 apply(drule mp)
  2653    apply(drule mp)
  2646 apply(simp add: CANDs_imply_SNa)
  2654     apply(simp add: CANDs_imply_SNa)
  2647 apply(drule_tac x="y" in spec)
  2655    apply(drule_tac x="y" in spec)
  2648 apply(rotate_tac 15)
  2656    apply(rotate_tac 15)
  2649 apply(drule mp)
  2657    apply(drule mp)
  2650 apply(assumption)
  2658     apply(assumption)
  2651 apply(rotate_tac 15)
  2659    apply(rotate_tac 15)
  2652 apply(drule mp)
  2660    apply(drule mp)
  2653 apply(simp add: CANDs_imply_SNa)
  2661     apply(simp add: CANDs_imply_SNa)
  2654 apply(assumption)
  2662    apply(assumption)
  2655 (* LAnd2 case *)
  2663     (* LAnd2 case *)
  2656 apply(erule disjE)
  2664   apply(erule disjE)
  2657 apply(erule exE)+
  2665    apply(erule exE)+
  2658 apply(auto)[1]
  2666    apply(auto)[1]
  2659 apply(frule_tac excluded_m)
  2667    apply(frule_tac excluded_m)
  2660 apply(assumption)
  2668     apply(assumption)
  2661 apply(erule disjE)
  2669    apply(erule disjE)
  2662 (* one of them in BINDING *)
  2670     (* one of them in BINDING *)
  2663 apply(erule disjE)
  2671     apply(erule disjE)
  2664 apply(drule fin_elims)
  2672      apply(drule fin_elims)
  2665 apply(drule fic_elims)
  2673      apply(drule fic_elims)
  2666 apply(simp)
  2674      apply(simp)
  2667 apply(drule BINDINGc_elim)
  2675      apply(drule BINDINGc_elim)
  2668 apply(drule_tac x="x" in spec)
  2676      apply(drule_tac x="x" in spec)
  2669 apply(drule_tac x="AndL2 (y).N' x" in spec)
  2677      apply(drule_tac x="AndL2 (y).N' x" in spec)
  2670 apply(simp)
  2678      apply(simp)
  2671 apply(simp add: better_AndR_substc)
  2679      apply(simp add: better_AndR_substc)
  2672 apply(generate_fresh "coname")
  2680      apply(generate_fresh "coname")
  2673 apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL2 (y).N' x) 
  2681      apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL2 (y).N' x) 
  2674                    = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x")
  2682                    = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x")
  2675 apply(simp)
  2683       apply(simp)
  2676 apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
  2684       apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
  2677 apply(auto intro: a_preserves_SNa)[1]
  2685        apply(auto intro: a_preserves_SNa)[1]
  2678 apply(rule al_redu)
  2686       apply(rule al_redu)
  2679 apply(rule better_LAnd2_intro)
  2687       apply(rule better_LAnd2_intro)
  2680 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2688        apply(simp add: abs_fresh fresh_prod fresh_atm)
  2681 apply(simp)
  2689       apply(simp)
  2682 apply(fresh_fun_simp (no_asm))
  2690      apply(fresh_fun_simp (no_asm))
  2683 apply(simp)
  2691      apply(simp)
  2684 (* other case of in BINDING *)
  2692     (* other case of in BINDING *)
  2685 apply(drule fin_elims)
  2693     apply(drule fin_elims)
  2686 apply(drule fic_elims)
  2694     apply(drule fic_elims)
  2687 apply(simp)
  2695     apply(simp)
  2688 apply(drule BINDINGn_elim)
  2696     apply(drule BINDINGn_elim)
  2689 apply(drule_tac x="a" in spec)
  2697     apply(drule_tac x="a" in spec)
  2690 apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
  2698     apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
  2691 apply(simp)
  2699     apply(simp)
  2692 apply(simp add: better_AndL2_substn)
  2700     apply(simp add: better_AndL2_substn)
  2693 apply(generate_fresh "name")
  2701     apply(generate_fresh "name")
  2694 apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL2 (y).N' z') 
  2702     apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL2 (y).N' z') 
  2695                    = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca")
  2703                    = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca")
  2696 apply(simp)
  2704      apply(simp)
  2697 apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
  2705      apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
  2698 apply(auto intro: a_preserves_SNa)[1]
  2706       apply(auto intro: a_preserves_SNa)[1]
  2699 apply(rule al_redu)
  2707      apply(rule al_redu)
  2700 apply(rule better_LAnd2_intro)
  2708      apply(rule better_LAnd2_intro)
  2701 apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2709       apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2702 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2710      apply(simp add: abs_fresh fresh_prod fresh_atm)
  2703 apply(fresh_fun_simp (no_asm))
  2711     apply(fresh_fun_simp (no_asm))
  2704 apply(simp)
  2712     apply(simp)
  2705 (* none of them in BINDING *)
  2713     (* none of them in BINDING *)
  2706 apply(simp)
  2714    apply(simp)
  2707 apply(erule conjE)
  2715    apply(erule conjE)
  2708 apply(frule CAND_AndR_elim)
  2716    apply(frule CAND_AndR_elim)
  2709 apply(assumption)
  2717     apply(assumption)
  2710 apply(erule exE)
  2718    apply(erule exE)
  2711 apply(frule CAND_AndL2_elim)
  2719    apply(frule CAND_AndL2_elim)
  2712 apply(assumption)
  2720     apply(assumption)
  2713 apply(erule exE)+
  2721    apply(erule exE)+
  2714 apply(simp only: ty.inject)
  2722    apply(simp only: ty.inject)
  2715 apply(drule_tac x="B2" in meta_spec)
  2723    apply(drule_tac x="B2" in meta_spec)
  2716 apply(drule_tac x="M2" in meta_spec)
  2724    apply(drule_tac x="M2" in meta_spec)
  2717 apply(drule_tac x="N'" in meta_spec)
  2725    apply(drule_tac x="N'" in meta_spec)
  2718 apply(erule conjE)+
  2726    apply(erule conjE)+
  2719 apply(drule mp)
  2727    apply(drule mp)
  2720 apply(simp)
  2728     apply(simp)
  2721 apply(drule_tac x="c" in spec)
  2729    apply(drule_tac x="c" in spec)
  2722 apply(rotate_tac 14)
  2730    apply(rotate_tac 14)
  2723 apply(drule mp)
  2731    apply(drule mp)
  2724 apply(assumption)
  2732     apply(assumption)
  2725 apply(rotate_tac 14)
  2733    apply(rotate_tac 14)
  2726 apply(drule mp)
  2734    apply(drule mp)
  2727 apply(simp add: CANDs_imply_SNa)
  2735     apply(simp add: CANDs_imply_SNa)
  2728 apply(drule_tac x="y" in spec)
  2736    apply(drule_tac x="y" in spec)
  2729 apply(rotate_tac 15)
  2737    apply(rotate_tac 15)
  2730 apply(drule mp)
  2738    apply(drule mp)
  2731 apply(assumption)
  2739     apply(assumption)
  2732 apply(rotate_tac 15)
  2740    apply(rotate_tac 15)
  2733 apply(drule mp)
  2741    apply(drule mp)
  2734 apply(simp add: CANDs_imply_SNa)
  2742     apply(simp add: CANDs_imply_SNa)
  2735 apply(assumption)
  2743    apply(assumption)
  2736 (* LOr1 case *)
  2744     (* LOr1 case *)
  2737 apply(erule disjE)
  2745   apply(erule disjE)
  2738 apply(erule exE)+
  2746    apply(erule exE)+
  2739 apply(auto)[1]
  2747    apply(auto)[1]
  2740 apply(frule_tac excluded_m)
  2748    apply(frule_tac excluded_m)
  2741 apply(assumption)
  2749     apply(assumption)
  2742 apply(erule disjE)
  2750    apply(erule disjE)
  2743 (* one of them in BINDING *)
  2751     (* one of them in BINDING *)
  2744 apply(erule disjE)
  2752     apply(erule disjE)
  2745 apply(drule fin_elims)
  2753      apply(drule fin_elims)
  2746 apply(drule fic_elims)
  2754      apply(drule fic_elims)
  2747 apply(simp)
  2755      apply(simp)
  2748 apply(drule BINDINGc_elim)
  2756      apply(drule BINDINGc_elim)
  2749 apply(drule_tac x="x" in spec)
  2757      apply(drule_tac x="x" in spec)
  2750 apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
  2758      apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
  2751 apply(simp)
  2759      apply(simp)
  2752 apply(simp add: better_OrR1_substc)
  2760      apply(simp add: better_OrR1_substc)
  2753 apply(generate_fresh "coname")
  2761      apply(generate_fresh "coname")
  2754 apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
  2762      apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
  2755                    = Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x")
  2763                    = Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x")
  2756 apply(simp)
  2764       apply(simp)
  2757 apply(subgoal_tac "Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
  2765       apply(subgoal_tac "Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
  2758 apply(auto intro: a_preserves_SNa)[1]
  2766        apply(auto intro: a_preserves_SNa)[1]
  2759 apply(rule al_redu)
  2767       apply(rule al_redu)
  2760 apply(rule better_LOr1_intro)
  2768       apply(rule better_LOr1_intro)
  2761 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2769        apply(simp add: abs_fresh fresh_prod fresh_atm)
  2762 apply(simp add: abs_fresh)
  2770       apply(simp add: abs_fresh)
  2763 apply(fresh_fun_simp (no_asm))
  2771      apply(fresh_fun_simp (no_asm))
  2764 apply(simp)
  2772      apply(simp)
  2765 (* other case of in BINDING *)
  2773     (* other case of in BINDING *)
  2766 apply(drule fin_elims)
  2774     apply(drule fin_elims)
  2767 apply(drule fic_elims)
  2775     apply(drule fic_elims)
  2768 apply(simp)
  2776     apply(simp)
  2769 apply(drule BINDINGn_elim)
  2777     apply(drule BINDINGn_elim)
  2770 apply(drule_tac x="a" in spec)
  2778     apply(drule_tac x="a" in spec)
  2771 apply(drule_tac x="OrR1 <b>.N' a" in spec)
  2779     apply(drule_tac x="OrR1 <b>.N' a" in spec)
  2772 apply(simp)
  2780     apply(simp)
  2773 apply(simp add: better_OrL_substn)
  2781     apply(simp add: better_OrL_substn)
  2774 apply(generate_fresh "name")
  2782     apply(generate_fresh "name")
  2775 apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR1 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
  2783     apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR1 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
  2776                    = Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c")
  2784                    = Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c")
  2777 apply(simp)
  2785      apply(simp)
  2778 apply(subgoal_tac "Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
  2786      apply(subgoal_tac "Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
  2779 apply(auto intro: a_preserves_SNa)[1]
  2787       apply(auto intro: a_preserves_SNa)[1]
  2780 apply(rule al_redu)
  2788      apply(rule al_redu)
  2781 apply(rule better_LOr1_intro)
  2789      apply(rule better_LOr1_intro)
  2782 apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2790       apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2783 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2791      apply(simp add: abs_fresh fresh_prod fresh_atm)
  2784 apply(fresh_fun_simp (no_asm))
  2792     apply(fresh_fun_simp (no_asm))
  2785 apply(simp)
  2793     apply(simp)
  2786 (* none of them in BINDING *)
  2794     (* none of them in BINDING *)
  2787 apply(simp)
  2795    apply(simp)
  2788 apply(erule conjE)
  2796    apply(erule conjE)
  2789 apply(frule CAND_OrR1_elim)
  2797    apply(frule CAND_OrR1_elim)
  2790 apply(assumption)
  2798     apply(assumption)
  2791 apply(erule exE)+
  2799    apply(erule exE)+
  2792 apply(frule CAND_OrL_elim)
  2800    apply(frule CAND_OrL_elim)
  2793 apply(assumption)
  2801     apply(assumption)
  2794 apply(erule exE)+
  2802    apply(erule exE)+
  2795 apply(simp only: ty.inject)
  2803    apply(simp only: ty.inject)
  2796 apply(drule_tac x="B1" in meta_spec)
  2804    apply(drule_tac x="B1" in meta_spec)
  2797 apply(drule_tac x="N'" in meta_spec)
  2805    apply(drule_tac x="N'" in meta_spec)
  2798 apply(drule_tac x="M1" in meta_spec)
  2806    apply(drule_tac x="M1" in meta_spec)
  2799 apply(erule conjE)+
  2807    apply(erule conjE)+
  2800 apply(drule mp)
  2808    apply(drule mp)
  2801 apply(simp)
  2809     apply(simp)
  2802 apply(drule_tac x="b" in spec)
  2810    apply(drule_tac x="b" in spec)
  2803 apply(rotate_tac 15)
  2811    apply(rotate_tac 15)
  2804 apply(drule mp)
  2812    apply(drule mp)
  2805 apply(assumption)
  2813     apply(assumption)
  2806 apply(rotate_tac 15)
  2814    apply(rotate_tac 15)
  2807 apply(drule mp)
  2815    apply(drule mp)
  2808 apply(simp add: CANDs_imply_SNa)
  2816     apply(simp add: CANDs_imply_SNa)
  2809 apply(drule_tac x="z" in spec)
  2817    apply(drule_tac x="z" in spec)
  2810 apply(rotate_tac 15)
  2818    apply(rotate_tac 15)
  2811 apply(drule mp)
  2819    apply(drule mp)
  2812 apply(assumption)
  2820     apply(assumption)
  2813 apply(rotate_tac 15)
  2821    apply(rotate_tac 15)
  2814 apply(drule mp)
  2822    apply(drule mp)
  2815 apply(simp add: CANDs_imply_SNa)
  2823     apply(simp add: CANDs_imply_SNa)
  2816 apply(assumption)
  2824    apply(assumption)
  2817 (* LOr2 case *)
  2825     (* LOr2 case *)
  2818 apply(erule disjE)
  2826   apply(erule disjE)
  2819 apply(erule exE)+
  2827    apply(erule exE)+
  2820 apply(auto)[1]
  2828    apply(auto)[1]
  2821 apply(frule_tac excluded_m)
  2829    apply(frule_tac excluded_m)
  2822 apply(assumption)
  2830     apply(assumption)
  2823 apply(erule disjE)
  2831    apply(erule disjE)
  2824 (* one of them in BINDING *)
  2832     (* one of them in BINDING *)
  2825 apply(erule disjE)
  2833     apply(erule disjE)
  2826 apply(drule fin_elims)
  2834      apply(drule fin_elims)
  2827 apply(drule fic_elims)
  2835      apply(drule fic_elims)
  2828 apply(simp)
  2836      apply(simp)
  2829 apply(drule BINDINGc_elim)
  2837      apply(drule BINDINGc_elim)
  2830 apply(drule_tac x="x" in spec)
  2838      apply(drule_tac x="x" in spec)
  2831 apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
  2839      apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
  2832 apply(simp)
  2840      apply(simp)
  2833 apply(simp add: better_OrR2_substc)
  2841      apply(simp add: better_OrR2_substc)
  2834 apply(generate_fresh "coname")
  2842      apply(generate_fresh "coname")
  2835 apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
  2843      apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
  2836                    = Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x")
  2844                    = Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x")
  2837 apply(simp)
  2845       apply(simp)
  2838 apply(subgoal_tac "Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
  2846       apply(subgoal_tac "Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
  2839 apply(auto intro: a_preserves_SNa)[1]
  2847        apply(auto intro: a_preserves_SNa)[1]
  2840 apply(rule al_redu)
  2848       apply(rule al_redu)
  2841 apply(rule better_LOr2_intro)
  2849       apply(rule better_LOr2_intro)
  2842 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2850        apply(simp add: abs_fresh fresh_prod fresh_atm)
  2843 apply(simp add: abs_fresh)
  2851       apply(simp add: abs_fresh)
  2844 apply(fresh_fun_simp (no_asm))
  2852      apply(fresh_fun_simp (no_asm))
  2845 apply(simp)
  2853      apply(simp)
  2846 (* other case of in BINDING *)
  2854     (* other case of in BINDING *)
  2847 apply(drule fin_elims)
  2855     apply(drule fin_elims)
  2848 apply(drule fic_elims)
  2856     apply(drule fic_elims)
  2849 apply(simp)
  2857     apply(simp)
  2850 apply(drule BINDINGn_elim)
  2858     apply(drule BINDINGn_elim)
  2851 apply(drule_tac x="a" in spec)
  2859     apply(drule_tac x="a" in spec)
  2852 apply(drule_tac x="OrR2 <b>.N' a" in spec)
  2860     apply(drule_tac x="OrR2 <b>.N' a" in spec)
  2853 apply(simp)
  2861     apply(simp)
  2854 apply(simp add: better_OrL_substn)
  2862     apply(simp add: better_OrL_substn)
  2855 apply(generate_fresh "name")
  2863     apply(generate_fresh "name")
  2856 apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR2 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
  2864     apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR2 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
  2857                    = Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c")
  2865                    = Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c")
  2858 apply(simp)
  2866      apply(simp)
  2859 apply(subgoal_tac "Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
  2867      apply(subgoal_tac "Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
  2860 apply(auto intro: a_preserves_SNa)[1]
  2868       apply(auto intro: a_preserves_SNa)[1]
  2861 apply(rule al_redu)
  2869      apply(rule al_redu)
  2862 apply(rule better_LOr2_intro)
  2870      apply(rule better_LOr2_intro)
  2863 apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2871       apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2864 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2872      apply(simp add: abs_fresh fresh_prod fresh_atm)
  2865 apply(fresh_fun_simp (no_asm))
  2873     apply(fresh_fun_simp (no_asm))
  2866 apply(simp)
  2874     apply(simp)
  2867 (* none of them in BINDING *)
  2875     (* none of them in BINDING *)
  2868 apply(simp)
  2876    apply(simp)
  2869 apply(erule conjE)
  2877    apply(erule conjE)
  2870 apply(frule CAND_OrR2_elim)
  2878    apply(frule CAND_OrR2_elim)
  2871 apply(assumption)
  2879     apply(assumption)
  2872 apply(erule exE)+
  2880    apply(erule exE)+
  2873 apply(frule CAND_OrL_elim)
  2881    apply(frule CAND_OrL_elim)
  2874 apply(assumption)
  2882     apply(assumption)
  2875 apply(erule exE)+
  2883    apply(erule exE)+
  2876 apply(simp only: ty.inject)
  2884    apply(simp only: ty.inject)
  2877 apply(drule_tac x="B2" in meta_spec)
  2885    apply(drule_tac x="B2" in meta_spec)
  2878 apply(drule_tac x="N'" in meta_spec)
  2886    apply(drule_tac x="N'" in meta_spec)
  2879 apply(drule_tac x="M2" in meta_spec)
  2887    apply(drule_tac x="M2" in meta_spec)
  2880 apply(erule conjE)+
  2888    apply(erule conjE)+
  2881 apply(drule mp)
  2889    apply(drule mp)
  2882 apply(simp)
  2890     apply(simp)
  2883 apply(drule_tac x="b" in spec)
  2891    apply(drule_tac x="b" in spec)
  2884 apply(rotate_tac 15)
  2892    apply(rotate_tac 15)
  2885 apply(drule mp)
  2893    apply(drule mp)
  2886 apply(assumption)
  2894     apply(assumption)
  2887 apply(rotate_tac 15)
  2895    apply(rotate_tac 15)
  2888 apply(drule mp)
  2896    apply(drule mp)
  2889 apply(simp add: CANDs_imply_SNa)
  2897     apply(simp add: CANDs_imply_SNa)
  2890 apply(drule_tac x="y" in spec)
  2898    apply(drule_tac x="y" in spec)
  2891 apply(rotate_tac 15)
  2899    apply(rotate_tac 15)
  2892 apply(drule mp)
  2900    apply(drule mp)
  2893 apply(assumption)
  2901     apply(assumption)
  2894 apply(rotate_tac 15)
  2902    apply(rotate_tac 15)
  2895 apply(drule mp)
  2903    apply(drule mp)
  2896 apply(simp add: CANDs_imply_SNa)
  2904     apply(simp add: CANDs_imply_SNa)
  2897 apply(assumption)
  2905    apply(assumption)
  2898 (* LImp case *)
  2906     (* LImp case *)
  2899 apply(erule exE)+
  2907   apply(erule exE)+
  2900 apply(auto)[1]
  2908   apply(auto)[1]
  2901 apply(frule_tac excluded_m)
  2909   apply(frule_tac excluded_m)
  2902 apply(assumption)
  2910    apply(assumption)
  2903 apply(erule disjE)
  2911   apply(erule disjE)
  2904 (* one of them in BINDING *)
  2912     (* one of them in BINDING *)
  2905 apply(erule disjE)
  2913    apply(erule disjE)
  2906 apply(drule fin_elims)
  2914     apply(drule fin_elims)
  2907 apply(drule fic_elims)
  2915     apply(drule fic_elims)
  2908 apply(simp)
  2916     apply(simp)
  2909 apply(drule BINDINGc_elim)
  2917     apply(drule BINDINGc_elim)
  2910 apply(drule_tac x="x" in spec)
  2918     apply(drule_tac x="x" in spec)
  2911 apply(drule_tac x="ImpL <c>.N1 (y).N2 x" in spec)
  2919     apply(drule_tac x="ImpL <c>.N1 (y).N2 x" in spec)
  2912 apply(simp)
  2920     apply(simp)
  2913 apply(simp add: better_ImpR_substc)
  2921     apply(simp add: better_ImpR_substc)
  2914 apply(generate_fresh "coname")
  2922     apply(generate_fresh "coname")
  2915 apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.ImpR (z).<b>.M'a a' (x).ImpL <c>.N1 (y).N2 x)
  2923     apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.ImpR (z).<b>.M'a a' (x).ImpL <c>.N1 (y).N2 x)
  2916                    = Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x")
  2924                    = Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x")
  2917 apply(simp)
  2925      apply(simp)
  2918 apply(subgoal_tac "Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x \<longrightarrow>\<^sub>a 
  2926      apply(subgoal_tac "Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x \<longrightarrow>\<^sub>a 
  2919                                                           Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
  2927                                                           Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
  2920 apply(auto intro: a_preserves_SNa)[1]
  2928       apply(auto intro: a_preserves_SNa)[1]
  2921 apply(rule al_redu)
  2929      apply(rule al_redu)
  2922 apply(rule better_LImp_intro)
  2930      apply(rule better_LImp_intro)
  2923 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2931        apply(simp add: abs_fresh fresh_prod fresh_atm)
  2924 apply(simp add: abs_fresh)
  2932       apply(simp add: abs_fresh)
  2925 apply(simp)
  2933      apply(simp)
  2926 apply(fresh_fun_simp (no_asm))
  2934     apply(fresh_fun_simp (no_asm))
  2927 apply(simp)
  2935     apply(simp)
  2928 (* other case of in BINDING *)
  2936     (* other case of in BINDING *)
  2929 apply(drule fin_elims)
  2937    apply(drule fin_elims)
  2930 apply(drule fic_elims)
  2938    apply(drule fic_elims)
  2931 apply(simp)
  2939    apply(simp)
  2932 apply(drule BINDINGn_elim)
  2940    apply(drule BINDINGn_elim)
  2933 apply(drule_tac x="a" in spec)
  2941    apply(drule_tac x="a" in spec)
  2934 apply(drule_tac x="ImpR (z).<b>.M'a a" in spec)
  2942    apply(drule_tac x="ImpR (z).<b>.M'a a" in spec)
  2935 apply(simp)
  2943    apply(simp)
  2936 apply(simp add: better_ImpL_substn)
  2944    apply(simp add: better_ImpL_substn)
  2937 apply(generate_fresh "name")
  2945    apply(generate_fresh "name")
  2938 apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.ImpR (z).<b>.M'a a (z').ImpL <c>.N1 (y).N2 z')
  2946    apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.ImpR (z).<b>.M'a a (z').ImpL <c>.N1 (y).N2 z')
  2939                    = Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca")
  2947                    = Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca")
  2940 apply(simp)
  2948     apply(simp)
  2941 apply(subgoal_tac "Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca \<longrightarrow>\<^sub>a 
  2949     apply(subgoal_tac "Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca \<longrightarrow>\<^sub>a 
  2942                                                           Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
  2950                                                           Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
  2943 apply(auto intro: a_preserves_SNa)[1]
  2951      apply(auto intro: a_preserves_SNa)[1]
  2944 apply(rule al_redu)
  2952     apply(rule al_redu)
  2945 apply(rule better_LImp_intro)
  2953     apply(rule better_LImp_intro)
  2946 apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2954       apply(simp add: abs_fresh fresh_prod fresh_atm) 
  2947 apply(simp add: abs_fresh fresh_prod fresh_atm)
  2955      apply(simp add: abs_fresh fresh_prod fresh_atm)
  2948 apply(simp)
  2956     apply(simp)
  2949 apply(fresh_fun_simp (no_asm))
  2957    apply(fresh_fun_simp (no_asm))
  2950 apply(simp add: abs_fresh abs_supp fin_supp)
  2958     apply(simp add: abs_fresh abs_supp fin_supp)
  2951 apply(simp add: abs_fresh abs_supp fin_supp)
  2959    apply(simp add: abs_fresh abs_supp fin_supp)
  2952 apply(simp)
  2960   apply(simp)
  2953 (* none of them in BINDING *)
  2961     (* none of them in BINDING *)
  2954 apply(erule conjE)
  2962   apply(erule conjE)
  2955 apply(frule CAND_ImpL_elim)
  2963   apply(frule CAND_ImpL_elim)
  2956 apply(assumption)
  2964    apply(assumption)
  2957 apply(erule exE)+
  2965   apply(erule exE)+
  2958 apply(frule CAND_ImpR_elim) (* check here *)
  2966   apply(frule CAND_ImpR_elim) (* check here *)
  2959 apply(assumption)
  2967    apply(assumption)
  2960 apply(erule exE)+
  2968   apply(erule exE)+
  2961 apply(erule conjE)+
  2969   apply(erule conjE)+
  2962 apply(simp only: ty.inject)
  2970   apply(simp only: ty.inject)
  2963 apply(erule conjE)+
  2971   apply(erule conjE)+
  2964 apply(case_tac "M'a=Ax z b")
  2972   apply(case_tac "M'a=Ax z b")
  2965 (* case Ma = Ax z b *)
  2973     (* case Ma = Ax z b *)
  2966 apply(rule_tac t="Cut <b>.(Cut <c>.N1 (z).M'a) (y).N2" and s="Cut <b>.(M'a{z:=<c>.N1}) (y).N2" in subst)
  2974    apply(rule_tac t="Cut <b>.(Cut <c>.N1 (z).M'a) (y).N2" and s="Cut <b>.(M'a{z:=<c>.N1}) (y).N2" in subst)
  2967 apply(simp)
  2975     apply(simp)
  2968 apply(drule_tac x="c" in spec)
  2976    apply(drule_tac x="c" in spec)
  2969 apply(drule_tac x="N1" in spec)
  2977    apply(drule_tac x="N1" in spec)
  2970 apply(drule mp)
  2978    apply(drule mp)
  2971 apply(simp)
  2979     apply(simp)
  2972 apply(drule_tac x="B2" in meta_spec)
  2980    apply(drule_tac x="B2" in meta_spec)
  2973 apply(drule_tac x="M'a{z:=<c>.N1}" in meta_spec)
  2981    apply(drule_tac x="M'a{z:=<c>.N1}" in meta_spec)
  2974 apply(drule_tac x="N2" in meta_spec)
  2982    apply(drule_tac x="N2" in meta_spec)
  2975 apply(drule conjunct1)
  2983    apply(drule conjunct1)
  2976 apply(drule mp)
  2984    apply(drule mp)
  2977 apply(simp)
  2985     apply(simp)
  2978 apply(rotate_tac 17)
  2986    apply(rotate_tac 17)
  2979 apply(drule_tac x="b" in spec)
  2987    apply(drule_tac x="b" in spec)
  2980 apply(drule mp)
  2988    apply(drule mp)
  2981 apply(assumption)
  2989     apply(assumption)
  2982 apply(drule mp)
  2990    apply(drule mp)
  2983 apply(simp add: CANDs_imply_SNa)
  2991     apply(simp add: CANDs_imply_SNa)
  2984 apply(rotate_tac 17)
  2992    apply(rotate_tac 17)
  2985 apply(drule_tac x="y" in spec)
  2993    apply(drule_tac x="y" in spec)
  2986 apply(drule mp)
  2994    apply(drule mp)
  2987 apply(assumption)
  2995     apply(assumption)
  2988 apply(drule mp)
  2996    apply(drule mp)
  2989 apply(simp add: CANDs_imply_SNa)
  2997     apply(simp add: CANDs_imply_SNa)
  2990 apply(assumption)
  2998    apply(assumption)
  2991 (* case Ma \<noteq> Ax z b *)
  2999     (* case Ma \<noteq> Ax z b *)
  2992 apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> \<parallel><B2>\<parallel>") (* lemma *)
  3000   apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> \<parallel><B2>\<parallel>") (* lemma *)
  2993 apply(frule_tac meta_spec)
  3001    apply(frule_tac meta_spec)
  2994 apply(drule_tac x="B2" in meta_spec)
  3002    apply(drule_tac x="B2" in meta_spec)
  2995 apply(drule_tac x="Cut <c>.N1 (z).M'a" in meta_spec)
  3003    apply(drule_tac x="Cut <c>.N1 (z).M'a" in meta_spec)
  2996 apply(drule_tac x="N2" in meta_spec)
  3004    apply(drule_tac x="N2" in meta_spec)
  2997 apply(erule conjE)+
  3005    apply(erule conjE)+
  2998 apply(drule mp)
  3006    apply(drule mp)
  2999 apply(simp)
  3007     apply(simp)
  3000 apply(rotate_tac 20)
  3008    apply(rotate_tac 20)
  3001 apply(drule_tac x="b" in spec)
  3009    apply(drule_tac x="b" in spec)
  3002 apply(rotate_tac 20)
  3010    apply(rotate_tac 20)
  3003 apply(drule mp)
  3011    apply(drule mp)
  3004 apply(assumption)
  3012     apply(assumption)
  3005 apply(rotate_tac 20)
  3013    apply(rotate_tac 20)
  3006 apply(drule mp)
  3014    apply(drule mp)
  3007 apply(simp add: CANDs_imply_SNa)
  3015     apply(simp add: CANDs_imply_SNa)
  3008 apply(rotate_tac 20)
  3016    apply(rotate_tac 20)
  3009 apply(drule_tac x="y" in spec)
  3017    apply(drule_tac x="y" in spec)
  3010 apply(rotate_tac 20)
  3018    apply(rotate_tac 20)
  3011 apply(drule mp)
  3019    apply(drule mp)
  3012 apply(assumption)
  3020     apply(assumption)
  3013 apply(rotate_tac 20)
  3021    apply(rotate_tac 20)
  3014 apply(drule mp)
  3022    apply(drule mp)
  3015 apply(simp add: CANDs_imply_SNa)
  3023     apply(simp add: CANDs_imply_SNa)
  3016 apply(assumption)
  3024    apply(assumption)
  3017 (* lemma *)
  3025     (* lemma *)
  3018 apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> BINDINGc B2 (\<parallel>(B2)\<parallel>)") (* second lemma *)
  3026   apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> BINDINGc B2 (\<parallel>(B2)\<parallel>)") (* second lemma *)
  3019 apply(simp add: BINDING_implies_CAND)
  3027    apply(simp add: BINDING_implies_CAND)
  3020 (* second lemma *)
  3028     (* second lemma *)
  3021 apply(simp (no_asm) add: BINDINGc_def)
  3029   apply(simp (no_asm) add: BINDINGc_def)
  3022 apply(rule exI)+
  3030   apply(rule exI)+
  3023 apply(rule conjI)
  3031   apply(rule conjI)
  3024 apply(rule refl)
  3032    apply(rule refl)
  3025 apply(rule allI)+
  3033   apply(rule allI)+
  3026 apply(rule impI)
  3034   apply(rule impI)
  3027 apply(generate_fresh "name")
  3035   apply(generate_fresh "name")
  3028 apply(rule_tac t="Cut <c>.N1 (z).M'a" and s="Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)" in subst)
  3036   apply(rule_tac t="Cut <c>.N1 (z).M'a" and s="Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)" in subst)
  3029 apply(simp add: trm.inject alpha fresh_prod fresh_atm)
  3037    apply(simp add: trm.inject alpha fresh_prod fresh_atm)
  3030 apply(rule_tac t="(Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)){b:=(xa).P}" 
  3038   apply(rule_tac t="(Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)){b:=(xa).P}" 
  3031            and s="Cut <c>.N1 (ca).(([(ca,z)]\<bullet>M'a){b:=(xa).P})" in subst)
  3039       and s="Cut <c>.N1 (ca).(([(ca,z)]\<bullet>M'a){b:=(xa).P})" in subst)
  3032 apply(rule sym)
  3040    apply(rule sym)
  3033 apply(rule tricky_subst)
  3041    apply(rule tricky_subst)
  3034 apply(simp)
  3042      apply(simp)
  3035 apply(simp)
  3043     apply(simp)
  3036 apply(clarify)
  3044    apply(clarify)
  3037 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
  3045    apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
  3038 apply(simp add: calc_atm)
  3046    apply(simp add: calc_atm)
  3039 apply(drule_tac x="B1" in meta_spec)
  3047   apply(drule_tac x="B1" in meta_spec)
  3040 apply(drule_tac x="N1" in meta_spec)
  3048   apply(drule_tac x="N1" in meta_spec)
  3041 apply(drule_tac x="([(ca,z)]\<bullet>M'a){b:=(xa).P}" in meta_spec)
  3049   apply(drule_tac x="([(ca,z)]\<bullet>M'a){b:=(xa).P}" in meta_spec)
  3042 apply(drule conjunct1)
  3050   apply(drule conjunct1)
  3043 apply(drule mp)
  3051   apply(drule mp)
  3044 apply(simp)
  3052    apply(simp)
  3045 apply(rotate_tac 19)
  3053   apply(rotate_tac 19)
  3046 apply(drule_tac x="c" in spec)
  3054   apply(drule_tac x="c" in spec)
  3047 apply(drule mp)
  3055   apply(drule mp)
  3048 apply(assumption)
  3056    apply(assumption)
  3049 apply(drule mp)
  3057   apply(drule mp)
  3050 apply(simp add: CANDs_imply_SNa)
  3058    apply(simp add: CANDs_imply_SNa)
  3051 apply(rotate_tac 19)
  3059   apply(rotate_tac 19)
  3052 apply(drule_tac x="ca" in spec)
  3060   apply(drule_tac x="ca" in spec)
  3053 apply(subgoal_tac "(ca):([(ca,z)]\<bullet>M'a){b:=(xa).P} \<in> \<parallel>(B1)\<parallel>")(*A*)
  3061   apply(subgoal_tac "(ca):([(ca,z)]\<bullet>M'a){b:=(xa).P} \<in> \<parallel>(B1)\<parallel>")(*A*)
  3054 apply(drule mp)
  3062    apply(drule mp)
  3055 apply(assumption)
  3063     apply(assumption)
  3056 apply(drule mp)
  3064    apply(drule mp)
  3057 apply(simp add: CANDs_imply_SNa)
  3065     apply(simp add: CANDs_imply_SNa)
  3058 apply(assumption)
  3066    apply(assumption)
  3059 (*A*)
  3067     (*A*)
  3060 apply(drule_tac x="[(ca,z)]\<bullet>xa" in spec)
  3068   apply(drule_tac x="[(ca,z)]\<bullet>xa" in spec)
  3061 apply(drule_tac x="[(ca,z)]\<bullet>P" in spec)
  3069   apply(drule_tac x="[(ca,z)]\<bullet>P" in spec)
  3062 apply(rotate_tac 19)
  3070   apply(rotate_tac 19)
  3063 apply(simp add: fresh_prod fresh_left)
  3071   apply(simp add: fresh_prod fresh_left)
  3064 apply(drule mp)
  3072   apply(drule mp)
  3065 apply(rule conjI)
  3073    apply(rule conjI)
  3066 apply(auto simp add: calc_atm)[1]
  3074     apply(auto simp add: calc_atm)[1]
  3067 apply(rule conjI)
  3075    apply(rule conjI)
  3068 apply(auto simp add: calc_atm)[1]
  3076     apply(auto simp add: calc_atm)[1]
  3069 apply(drule_tac pi="[(ca,z)]" and x="(xa):P" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
  3077    apply(drule_tac pi="[(ca,z)]" and x="(xa):P" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
  3070 apply(simp add: CAND_eqvt_name)
  3078    apply(simp add: CAND_eqvt_name)
  3071 apply(drule_tac pi="[(ca,z)]" and X="\<parallel>(B1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
  3079   apply(drule_tac pi="[(ca,z)]" and X="\<parallel>(B1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
  3072 apply(simp add: CAND_eqvt_name csubst_eqvt)
  3080   apply(simp add: CAND_eqvt_name csubst_eqvt)
  3073 apply(perm_simp)
  3081   apply(perm_simp)
  3074 done
  3082   done
  3075 
  3083 
  3076 
  3084 
  3077 (* parallel substitution *)
  3085 (* parallel substitution *)
  3078 
  3086 
  3079 
  3087 
  3080 lemma CUT_SNa:
  3088 lemma CUT_SNa:
  3081   assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
  3089   assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
  3082   and     a2: "(x):N \<in> (\<parallel>(B)\<parallel>)"
  3090     and     a2: "(x):N \<in> (\<parallel>(B)\<parallel>)"
  3083   shows   "SNa (Cut <a>.M (x).N)"
  3091   shows   "SNa (Cut <a>.M (x).N)"
  3084 using a1 a2
  3092   using a1 a2
  3085 apply -
  3093   apply -
  3086 apply(rule CUT_SNa_aux[OF a1])
  3094   apply(rule CUT_SNa_aux[OF a1])
  3087 apply(simp_all add: CANDs_imply_SNa)
  3095     apply(simp_all add: CANDs_imply_SNa)
  3088 done 
  3096   done 
  3089 
  3097 
  3090 
  3098 
  3091 fun 
  3099 fun 
  3092  findn :: "(name\<times>coname\<times>trm) list\<Rightarrow>name\<Rightarrow>(coname\<times>trm) option"
  3100   findn :: "(name\<times>coname\<times>trm) list\<Rightarrow>name\<Rightarrow>(coname\<times>trm) option"
  3093 where
  3101   where
  3094   "findn [] x = None"
  3102     "findn [] x = None"
  3095 | "findn ((y,c,P)#\<theta>_n) x = (if y=x then Some (c,P) else findn \<theta>_n x)"
  3103   | "findn ((y,c,P)#\<theta>_n) x = (if y=x then Some (c,P) else findn \<theta>_n x)"
  3096 
  3104 
  3097 lemma findn_eqvt[eqvt]:
  3105 lemma findn_eqvt[eqvt]:
  3098   fixes pi1::"name prm"
  3106   fixes pi1::"name prm"
  3099   and   pi2::"coname prm"
  3107     and   pi2::"coname prm"
  3100   shows "(pi1\<bullet>findn \<theta>_n x) = findn (pi1\<bullet>\<theta>_n) (pi1\<bullet>x)" 
  3108   shows "(pi1\<bullet>findn \<theta>_n x) = findn (pi1\<bullet>\<theta>_n) (pi1\<bullet>x)" 
  3101   and   "(pi2\<bullet>findn \<theta>_n x) = findn (pi2\<bullet>\<theta>_n) (pi2\<bullet>x)"
  3109     and   "(pi2\<bullet>findn \<theta>_n x) = findn (pi2\<bullet>\<theta>_n) (pi2\<bullet>x)"
  3102 apply(induct \<theta>_n)
  3110    apply(induct \<theta>_n)
  3103 apply(auto simp add: perm_bij) 
  3111      apply(auto simp add: perm_bij) 
  3104 done
  3112   done
  3105 
  3113 
  3106 lemma findn_fresh:
  3114 lemma findn_fresh:
  3107   assumes a: "x\<sharp>\<theta>_n"
  3115   assumes a: "x\<sharp>\<theta>_n"
  3108   shows "findn \<theta>_n x = None"
  3116   shows "findn \<theta>_n x = None"
  3109 using a
  3117   using a
  3110 apply(induct \<theta>_n)
  3118   apply(induct \<theta>_n)
  3111 apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
  3119    apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
  3112 done
  3120   done
  3113 
  3121 
  3114 fun 
  3122 fun 
  3115  findc :: "(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>(name\<times>trm) option"
  3123   findc :: "(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>(name\<times>trm) option"
  3116 where
  3124   where
  3117   "findc [] x = None"
  3125     "findc [] x = None"
  3118 | "findc ((c,y,P)#\<theta>_c) a = (if a=c then Some (y,P) else findc \<theta>_c a)"
  3126   | "findc ((c,y,P)#\<theta>_c) a = (if a=c then Some (y,P) else findc \<theta>_c a)"
  3119 
  3127 
  3120 lemma findc_eqvt[eqvt]:
  3128 lemma findc_eqvt[eqvt]:
  3121   fixes pi1::"name prm"
  3129   fixes pi1::"name prm"
  3122   and   pi2::"coname prm"
  3130     and   pi2::"coname prm"
  3123   shows "(pi1\<bullet>findc \<theta>_c a) = findc (pi1\<bullet>\<theta>_c) (pi1\<bullet>a)" 
  3131   shows "(pi1\<bullet>findc \<theta>_c a) = findc (pi1\<bullet>\<theta>_c) (pi1\<bullet>a)" 
  3124   and   "(pi2\<bullet>findc \<theta>_c a) = findc (pi2\<bullet>\<theta>_c) (pi2\<bullet>a)"
  3132     and   "(pi2\<bullet>findc \<theta>_c a) = findc (pi2\<bullet>\<theta>_c) (pi2\<bullet>a)"
  3125 apply(induct \<theta>_c)
  3133    apply(induct \<theta>_c)
  3126 apply(auto simp add: perm_bij) 
  3134      apply(auto simp add: perm_bij) 
  3127 done
  3135   done
  3128 
  3136 
  3129 lemma findc_fresh:
  3137 lemma findc_fresh:
  3130   assumes a: "a\<sharp>\<theta>_c"
  3138   assumes a: "a\<sharp>\<theta>_c"
  3131   shows "findc \<theta>_c a = None"
  3139   shows "findc \<theta>_c a = None"
  3132 using a
  3140   using a
  3133 apply(induct \<theta>_c)
  3141   apply(induct \<theta>_c)
  3134 apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
  3142    apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
  3135 done
  3143   done
  3136 
  3144 
  3137 abbreviation 
  3145 abbreviation 
  3138  nmaps :: "(name\<times>coname\<times>trm) list \<Rightarrow> name \<Rightarrow> (coname\<times>trm) option \<Rightarrow> bool" ("_ nmaps _ to _" [55,55,55] 55) 
  3146   nmaps :: "(name\<times>coname\<times>trm) list \<Rightarrow> name \<Rightarrow> (coname\<times>trm) option \<Rightarrow> bool" ("_ nmaps _ to _" [55,55,55] 55) 
  3139 where
  3147   where
  3140  "\<theta>_n nmaps x to P \<equiv> (findn \<theta>_n x) = P"
  3148     "\<theta>_n nmaps x to P \<equiv> (findn \<theta>_n x) = P"
  3141 
  3149 
  3142 abbreviation 
  3150 abbreviation 
  3143  cmaps :: "(coname\<times>name\<times>trm) list \<Rightarrow> coname \<Rightarrow> (name\<times>trm) option \<Rightarrow> bool" ("_ cmaps _ to _" [55,55,55] 55) 
  3151   cmaps :: "(coname\<times>name\<times>trm) list \<Rightarrow> coname \<Rightarrow> (name\<times>trm) option \<Rightarrow> bool" ("_ cmaps _ to _" [55,55,55] 55) 
  3144 where
  3152   where
  3145  "\<theta>_c cmaps a to P \<equiv> (findc \<theta>_c a) = P"
  3153     "\<theta>_c cmaps a to P \<equiv> (findc \<theta>_c a) = P"
  3146 
  3154 
  3147 lemma nmaps_fresh:
  3155 lemma nmaps_fresh:
  3148   shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> a\<sharp>\<theta>_n \<Longrightarrow> a\<sharp>(c,P)"
  3156   shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> a\<sharp>\<theta>_n \<Longrightarrow> a\<sharp>(c,P)"
  3149 apply(induct \<theta>_n)
  3157   apply(induct \<theta>_n)
  3150 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3158    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3151 apply(case_tac "aa=x")
  3159    apply(case_tac "aa=x")
  3152 apply(auto)
  3160     apply(auto)
  3153 apply(case_tac "aa=x")
  3161   apply(case_tac "aa=x")
  3154 apply(auto)
  3162    apply(auto)
  3155 done
  3163   done
  3156 
  3164 
  3157 lemma cmaps_fresh:
  3165 lemma cmaps_fresh:
  3158   shows "\<theta>_c cmaps a to Some (y,P) \<Longrightarrow> x\<sharp>\<theta>_c \<Longrightarrow> x\<sharp>(y,P)"
  3166   shows "\<theta>_c cmaps a to Some (y,P) \<Longrightarrow> x\<sharp>\<theta>_c \<Longrightarrow> x\<sharp>(y,P)"
  3159 apply(induct \<theta>_c)
  3167   apply(induct \<theta>_c)
  3160 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3168    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3161 apply(case_tac "a=aa")
  3169    apply(case_tac "a=aa")
  3162 apply(auto)
  3170     apply(auto)
  3163 apply(case_tac "a=aa")
  3171   apply(case_tac "a=aa")
  3164 apply(auto)
  3172    apply(auto)
  3165 done
  3173   done
  3166 
  3174 
  3167 lemma nmaps_false:
  3175 lemma nmaps_false:
  3168   shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> x\<sharp>\<theta>_n \<Longrightarrow> False"
  3176   shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> x\<sharp>\<theta>_n \<Longrightarrow> False"
  3169 apply(induct \<theta>_n)
  3177   apply(induct \<theta>_n)
  3170 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3178    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3171 done
  3179   done
  3172 
  3180 
  3173 lemma cmaps_false:
  3181 lemma cmaps_false:
  3174   shows "\<theta>_c cmaps c to Some (x,P) \<Longrightarrow> c\<sharp>\<theta>_c \<Longrightarrow> False"
  3182   shows "\<theta>_c cmaps c to Some (x,P) \<Longrightarrow> c\<sharp>\<theta>_c \<Longrightarrow> False"
  3175 apply(induct \<theta>_c)
  3183   apply(induct \<theta>_c)
  3176 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3184    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3177 done
  3185   done
  3178 
  3186 
  3179 fun 
  3187 fun 
  3180  lookupa :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
  3188   lookupa :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
  3181 where
  3189   where
  3182   "lookupa x a [] = Ax x a"
  3190     "lookupa x a [] = Ax x a"
  3183 | "lookupa x a ((c,y,P)#\<theta>_c) = (if a=c then Cut <a>.Ax x a (y).P else lookupa x a \<theta>_c)"
  3191   | "lookupa x a ((c,y,P)#\<theta>_c) = (if a=c then Cut <a>.Ax x a (y).P else lookupa x a \<theta>_c)"
  3184 
  3192 
  3185 lemma lookupa_eqvt[eqvt]:
  3193 lemma lookupa_eqvt[eqvt]:
  3186   fixes pi1::"name prm"
  3194   fixes pi1::"name prm"
  3187   and   pi2::"coname prm"
  3195     and   pi2::"coname prm"
  3188   shows "(pi1\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_c)"
  3196   shows "(pi1\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_c)"
  3189   and   "(pi2\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_c)"
  3197     and   "(pi2\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_c)"
  3190 apply -
  3198    apply -
  3191 apply(induct \<theta>_c)
  3199    apply(induct \<theta>_c)
  3192 apply(auto simp add: eqvts)
  3200     apply(auto simp add: eqvts)
  3193 apply(induct \<theta>_c)
  3201   apply(induct \<theta>_c)
  3194 apply(auto simp add: eqvts)
  3202    apply(auto simp add: eqvts)
  3195 done
  3203   done
  3196 
  3204 
  3197 lemma lookupa_fire:
  3205 lemma lookupa_fire:
  3198   assumes a: "\<theta>_c cmaps a to Some (y,P)"
  3206   assumes a: "\<theta>_c cmaps a to Some (y,P)"
  3199   shows "(lookupa x a \<theta>_c) = Cut <a>.Ax x a (y).P"
  3207   shows "(lookupa x a \<theta>_c) = Cut <a>.Ax x a (y).P"
  3200 using a
  3208   using a
  3201 apply(induct \<theta>_c arbitrary: x a y P)
  3209   apply(induct \<theta>_c arbitrary: x a y P)
  3202 apply(auto)
  3210    apply(auto)
  3203 done
  3211   done
  3204 
  3212 
  3205 fun 
  3213 fun 
  3206  lookupb :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>trm\<Rightarrow>trm"
  3214   lookupb :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>trm\<Rightarrow>trm"
  3207 where
  3215   where
  3208   "lookupb x a [] c P = Cut <c>.P (x).Ax x a"
  3216     "lookupb x a [] c P = Cut <c>.P (x).Ax x a"
  3209 | "lookupb x a ((d,y,N)#\<theta>_c) c P = (if a=d then Cut <c>.P (y).N  else lookupb x a \<theta>_c c P)"
  3217   | "lookupb x a ((d,y,N)#\<theta>_c) c P = (if a=d then Cut <c>.P (y).N  else lookupb x a \<theta>_c c P)"
  3210 
  3218 
  3211 lemma lookupb_eqvt[eqvt]:
  3219 lemma lookupb_eqvt[eqvt]:
  3212   fixes pi1::"name prm"
  3220   fixes pi1::"name prm"
  3213   and   pi2::"coname prm"
  3221     and   pi2::"coname prm"
  3214   shows "(pi1\<bullet>(lookupb x a \<theta>_c c P)) = lookupb (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_c) (pi1\<bullet>c) (pi1\<bullet>P)"
  3222   shows "(pi1\<bullet>(lookupb x a \<theta>_c c P)) = lookupb (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_c) (pi1\<bullet>c) (pi1\<bullet>P)"
  3215   and   "(pi2\<bullet>(lookupb x a \<theta>_c c P)) = lookupb (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_c) (pi2\<bullet>c) (pi2\<bullet>P)"
  3223     and   "(pi2\<bullet>(lookupb x a \<theta>_c c P)) = lookupb (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_c) (pi2\<bullet>c) (pi2\<bullet>P)"
  3216 apply -
  3224    apply -
  3217 apply(induct \<theta>_c)
  3225    apply(induct \<theta>_c)
  3218 apply(auto simp add: eqvts)
  3226     apply(auto simp add: eqvts)
  3219 apply(induct \<theta>_c)
  3227   apply(induct \<theta>_c)
  3220 apply(auto simp add: eqvts)
  3228    apply(auto simp add: eqvts)
  3221 done
  3229   done
  3222 
  3230 
  3223 fun 
  3231 fun 
  3224   lookup :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
  3232   lookup :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
  3225 where
  3233   where
  3226   "lookup x a [] \<theta>_c = lookupa x a \<theta>_c"
  3234     "lookup x a [] \<theta>_c = lookupa x a \<theta>_c"
  3227 | "lookup x a ((y,c,P)#\<theta>_n) \<theta>_c = (if x=y then (lookupb x a \<theta>_c c P) else lookup x a \<theta>_n \<theta>_c)"
  3235   | "lookup x a ((y,c,P)#\<theta>_n) \<theta>_c = (if x=y then (lookupb x a \<theta>_c c P) else lookup x a \<theta>_n \<theta>_c)"
  3228 
  3236 
  3229 lemma lookup_eqvt[eqvt]:
  3237 lemma lookup_eqvt[eqvt]:
  3230   fixes pi1::"name prm"
  3238   fixes pi1::"name prm"
  3231   and   pi2::"coname prm"
  3239     and   pi2::"coname prm"
  3232   shows "(pi1\<bullet>(lookup x a \<theta>_n \<theta>_c)) = lookup (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n) (pi1\<bullet>\<theta>_c)"
  3240   shows "(pi1\<bullet>(lookup x a \<theta>_n \<theta>_c)) = lookup (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n) (pi1\<bullet>\<theta>_c)"
  3233   and   "(pi2\<bullet>(lookup x a \<theta>_n \<theta>_c)) = lookup (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n) (pi2\<bullet>\<theta>_c)"
  3241     and   "(pi2\<bullet>(lookup x a \<theta>_n \<theta>_c)) = lookup (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n) (pi2\<bullet>\<theta>_c)"
  3234 apply -
  3242    apply -
  3235 apply(induct \<theta>_n)
  3243    apply(induct \<theta>_n)
  3236 apply(auto simp add: eqvts)
  3244     apply(auto simp add: eqvts)
  3237 apply(induct \<theta>_n)
  3245   apply(induct \<theta>_n)
  3238 apply(auto simp add: eqvts)
  3246    apply(auto simp add: eqvts)
  3239 done
  3247   done
  3240 
  3248 
  3241 fun 
  3249 fun 
  3242   lookupc :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm"
  3250   lookupc :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm"
  3243 where
  3251   where
  3244   "lookupc x a [] = Ax x a"
  3252     "lookupc x a [] = Ax x a"
  3245 | "lookupc x a ((y,c,P)#\<theta>_n) = (if x=y then P[c\<turnstile>c>a] else lookupc x a \<theta>_n)"
  3253   | "lookupc x a ((y,c,P)#\<theta>_n) = (if x=y then P[c\<turnstile>c>a] else lookupc x a \<theta>_n)"
  3246 
  3254 
  3247 lemma lookupc_eqvt[eqvt]:
  3255 lemma lookupc_eqvt[eqvt]:
  3248   fixes pi1::"name prm"
  3256   fixes pi1::"name prm"
  3249   and   pi2::"coname prm"
  3257     and   pi2::"coname prm"
  3250   shows "(pi1\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
  3258   shows "(pi1\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
  3251   and   "(pi2\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
  3259     and   "(pi2\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
  3252 apply -
  3260    apply -
  3253 apply(induct \<theta>_n)
  3261    apply(induct \<theta>_n)
  3254 apply(auto simp add: eqvts)
  3262     apply(auto simp add: eqvts)
  3255 apply(induct \<theta>_n)
  3263   apply(induct \<theta>_n)
  3256 apply(auto simp add: eqvts)
  3264    apply(auto simp add: eqvts)
  3257 done
  3265   done
  3258 
  3266 
  3259 fun 
  3267 fun 
  3260   lookupd :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
  3268   lookupd :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
  3261 where
  3269   where
  3262   "lookupd x a [] = Ax x a"
  3270     "lookupd x a [] = Ax x a"
  3263 | "lookupd x a ((c,y,P)#\<theta>_c) = (if a=c then P[y\<turnstile>n>x] else lookupd x a \<theta>_c)"
  3271   | "lookupd x a ((c,y,P)#\<theta>_c) = (if a=c then P[y\<turnstile>n>x] else lookupd x a \<theta>_c)"
  3264 
  3272 
  3265 lemma lookupd_eqvt[eqvt]:
  3273 lemma lookupd_eqvt[eqvt]:
  3266   fixes pi1::"name prm"
  3274   fixes pi1::"name prm"
  3267   and   pi2::"coname prm"
  3275     and   pi2::"coname prm"
  3268   shows "(pi1\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
  3276   shows "(pi1\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
  3269   and   "(pi2\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
  3277     and   "(pi2\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
  3270 apply -
  3278    apply -
  3271 apply(induct \<theta>_n)
  3279    apply(induct \<theta>_n)
  3272 apply(auto simp add: eqvts)
  3280     apply(auto simp add: eqvts)
  3273 apply(induct \<theta>_n)
  3281   apply(induct \<theta>_n)
  3274 apply(auto simp add: eqvts)
  3282    apply(auto simp add: eqvts)
  3275 done
  3283   done
  3276 
  3284 
  3277 lemma lookupa_fresh:
  3285 lemma lookupa_fresh:
  3278   assumes a: "a\<sharp>\<theta>_c"
  3286   assumes a: "a\<sharp>\<theta>_c"
  3279   shows "lookupa y a \<theta>_c = Ax y a"
  3287   shows "lookupa y a \<theta>_c = Ax y a"
  3280 using a
  3288   using a
  3281 apply(induct \<theta>_c)
  3289   apply(induct \<theta>_c)
  3282 apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3290    apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3283 done
  3291   done
  3284 
  3292 
  3285 lemma lookupa_csubst:
  3293 lemma lookupa_csubst:
  3286   assumes a: "a\<sharp>\<theta>_c"
  3294   assumes a: "a\<sharp>\<theta>_c"
  3287   shows "Cut <a>.Ax y a (x).P = (lookupa y a \<theta>_c){a:=(x).P}"
  3295   shows "Cut <a>.Ax y a (x).P = (lookupa y a \<theta>_c){a:=(x).P}"
  3288 using a by (simp add: lookupa_fresh)
  3296   using a by (simp add: lookupa_fresh)
  3289 
  3297 
  3290 lemma lookupa_freshness:
  3298 lemma lookupa_freshness:
  3291   fixes a::"coname"
  3299   fixes a::"coname"
  3292   and   x::"name"
  3300     and   x::"name"
  3293   shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupa y c \<theta>_c"
  3301   shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupa y c \<theta>_c"
  3294   and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupa y c \<theta>_c"
  3302     and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupa y c \<theta>_c"
  3295 apply(induct \<theta>_c)
  3303    apply(induct \<theta>_c)
  3296 apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3304      apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3297 done
  3305   done
  3298 
  3306 
  3299 lemma lookupa_unicity:
  3307 lemma lookupa_unicity:
  3300   assumes a: "lookupa x a \<theta>_c= Ax y b" "b\<sharp>\<theta>_c" "y\<sharp>\<theta>_c"
  3308   assumes a: "lookupa x a \<theta>_c= Ax y b" "b\<sharp>\<theta>_c" "y\<sharp>\<theta>_c"
  3301   shows "x=y \<and> a=b"
  3309   shows "x=y \<and> a=b"
  3302 using a
  3310   using a
  3303 apply(induct \<theta>_c)
  3311   apply(induct \<theta>_c)
  3304 apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
  3312    apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
  3305 apply(case_tac "a=aa")
  3313    apply(case_tac "a=aa")
  3306 apply(auto)
  3314     apply(auto)
  3307 apply(case_tac "a=aa")
  3315   apply(case_tac "a=aa")
  3308 apply(auto)
  3316    apply(auto)
  3309 done
  3317   done
  3310 
  3318 
  3311 lemma lookupb_csubst:
  3319 lemma lookupb_csubst:
  3312   assumes a: "a\<sharp>(\<theta>_c,c,N)"
  3320   assumes a: "a\<sharp>(\<theta>_c,c,N)"
  3313   shows "Cut <c>.N (x).P = (lookupb y a \<theta>_c c N){a:=(x).P}"
  3321   shows "Cut <c>.N (x).P = (lookupb y a \<theta>_c c N){a:=(x).P}"
  3314 using a
  3322   using a
  3315 apply(induct \<theta>_c)
  3323   apply(induct \<theta>_c)
  3316 apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
  3324    apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
  3317 apply(rule sym)
  3325   apply(rule sym)
  3318 apply(generate_fresh "name")
  3326   apply(generate_fresh "name")
  3319 apply(generate_fresh "coname")
  3327   apply(generate_fresh "coname")
  3320 apply(subgoal_tac "Cut <c>.N (y).Ax y a = Cut <caa>.([(caa,c)]\<bullet>N) (ca).Ax ca a")
  3328   apply(subgoal_tac "Cut <c>.N (y).Ax y a = Cut <caa>.([(caa,c)]\<bullet>N) (ca).Ax ca a")
  3321 apply(simp)
  3329    apply(simp)
  3322 apply(rule trans)
  3330    apply(rule trans)
  3323 apply(rule better_Cut_substc)
  3331     apply(rule better_Cut_substc)
  3324 apply(simp)
  3332      apply(simp)
  3325 apply(simp add: abs_fresh)
  3333     apply(simp add: abs_fresh)
  3326 apply(simp)
  3334    apply(simp)
  3327 apply(subgoal_tac "a\<sharp>([(caa,c)]\<bullet>N)")
  3335    apply(subgoal_tac "a\<sharp>([(caa,c)]\<bullet>N)")
  3328 apply(simp add: forget)
  3336     apply(simp add: forget)
  3329 apply(simp add: trm.inject)
  3337     apply(simp add: trm.inject)
  3330 apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
  3338    apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
  3331 apply(simp add: trm.inject)
  3339   apply(simp add: trm.inject)
  3332 apply(rule conjI)
  3340   apply(rule conjI)
  3333 apply(rule sym)
  3341    apply(rule sym)
  3334 apply(simp add: alpha fresh_prod fresh_atm)
  3342    apply(simp add: alpha fresh_prod fresh_atm)
  3335 apply(simp add: alpha calc_atm fresh_prod fresh_atm)
  3343   apply(simp add: alpha calc_atm fresh_prod fresh_atm)
  3336 done
  3344   done
  3337 
  3345 
  3338 lemma lookupb_freshness:
  3346 lemma lookupb_freshness:
  3339   fixes a::"coname"
  3347   fixes a::"coname"
  3340   and   x::"name"
  3348     and   x::"name"
  3341   shows "a\<sharp>(\<theta>_c,c,b,P) \<Longrightarrow> a\<sharp>lookupb y c \<theta>_c b P"
  3349   shows "a\<sharp>(\<theta>_c,c,b,P) \<Longrightarrow> a\<sharp>lookupb y c \<theta>_c b P"
  3342   and   "x\<sharp>(\<theta>_c,y,P) \<Longrightarrow> x\<sharp>lookupb y c \<theta>_c b P"
  3350     and   "x\<sharp>(\<theta>_c,y,P) \<Longrightarrow> x\<sharp>lookupb y c \<theta>_c b P"
  3343 apply(induct \<theta>_c)
  3351    apply(induct \<theta>_c)
  3344 apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3352      apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3345 done
  3353   done
  3346 
  3354 
  3347 lemma lookupb_unicity:
  3355 lemma lookupb_unicity:
  3348   assumes a: "lookupb x a \<theta>_c c P = Ax y b" "b\<sharp>(\<theta>_c,c,P)" "y\<sharp>\<theta>_c"
  3356   assumes a: "lookupb x a \<theta>_c c P = Ax y b" "b\<sharp>(\<theta>_c,c,P)" "y\<sharp>\<theta>_c"
  3349   shows "x=y \<and> a=b"
  3357   shows "x=y \<and> a=b"
  3350 using a
  3358   using a
  3351 apply(induct \<theta>_c)
  3359   apply(induct \<theta>_c)
  3352 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3360    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  3353 apply(case_tac "a=aa")
  3361    apply(case_tac "a=aa")
  3354 apply(auto)
  3362     apply(auto)
  3355 apply(case_tac "a=aa")
  3363   apply(case_tac "a=aa")
  3356 apply(auto)
  3364    apply(auto)
  3357 done
  3365   done
  3358 
  3366 
  3359 lemma lookupb_lookupa:
  3367 lemma lookupb_lookupa:
  3360   assumes a: "x\<sharp>\<theta>_c"
  3368   assumes a: "x\<sharp>\<theta>_c"
  3361   shows "lookupb x c \<theta>_c a P = (lookupa x c \<theta>_c){x:=<a>.P}"
  3369   shows "lookupb x c \<theta>_c a P = (lookupa x c \<theta>_c){x:=<a>.P}"
  3362 using a
  3370   using a
  3363 apply(induct \<theta>_c)
  3371   apply(induct \<theta>_c)
  3364 apply(auto simp add: fresh_list_cons fresh_prod)
  3372    apply(auto simp add: fresh_list_cons fresh_prod)
  3365 apply(generate_fresh "coname")
  3373   apply(generate_fresh "coname")
  3366 apply(generate_fresh "name")
  3374   apply(generate_fresh "name")
  3367 apply(subgoal_tac "Cut <c>.Ax x c (aa).b = Cut <ca>.Ax x ca (caa).([(caa,aa)]\<bullet>b)")
  3375   apply(subgoal_tac "Cut <c>.Ax x c (aa).b = Cut <ca>.Ax x ca (caa).([(caa,aa)]\<bullet>b)")
  3368 apply(simp)
  3376    apply(simp)
  3369 apply(rule sym)
  3377    apply(rule sym)
  3370 apply(rule trans)
  3378    apply(rule trans)
  3371 apply(rule better_Cut_substn)
  3379     apply(rule better_Cut_substn)
  3372 apply(simp add: abs_fresh)
  3380      apply(simp add: abs_fresh)
  3373 apply(simp)
  3381     apply(simp)
  3374 apply(simp)
  3382    apply(simp)
  3375 apply(subgoal_tac "x\<sharp>([(caa,aa)]\<bullet>b)")
  3383    apply(subgoal_tac "x\<sharp>([(caa,aa)]\<bullet>b)")
  3376 apply(simp add: forget)
  3384     apply(simp add: forget)
  3377 apply(simp add: trm.inject)
  3385     apply(simp add: trm.inject)
  3378 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
  3386    apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
  3379 apply(simp add: trm.inject)
  3387   apply(simp add: trm.inject)
  3380 apply(rule conjI)
  3388   apply(rule conjI)
  3381 apply(simp add: alpha calc_atm fresh_atm fresh_prod)
  3389    apply(simp add: alpha calc_atm fresh_atm fresh_prod)
  3382 apply(rule sym)
  3390   apply(rule sym)
  3383 apply(simp add: alpha calc_atm fresh_atm fresh_prod)
  3391   apply(simp add: alpha calc_atm fresh_atm fresh_prod)
  3384 done
  3392   done
  3385 
  3393 
  3386 lemma lookup_csubst:
  3394 lemma lookup_csubst:
  3387   assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
  3395   assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
  3388   shows "lookup y c \<theta>_n ((a,x,P)#\<theta>_c) = (lookup y c \<theta>_n \<theta>_c){a:=(x).P}"
  3396   shows "lookup y c \<theta>_n ((a,x,P)#\<theta>_c) = (lookup y c \<theta>_n \<theta>_c){a:=(x).P}"
  3389 using a
  3397   using a
  3390 apply(induct \<theta>_n)
  3398   apply(induct \<theta>_n)
  3391 apply(auto simp add: fresh_prod fresh_list_cons)
  3399    apply(auto simp add: fresh_prod fresh_list_cons)
  3392 apply(simp add: lookupa_csubst)
  3400      apply(simp add: lookupa_csubst)
  3393 apply(simp add: lookupa_freshness forget fresh_atm fresh_prod)
  3401     apply(simp add: lookupa_freshness forget fresh_atm fresh_prod)
  3394 apply(rule lookupb_csubst)
  3402    apply(rule lookupb_csubst)
  3395 apply(simp)
  3403    apply(simp)
  3396 apply(auto simp add: lookupb_freshness forget fresh_atm fresh_prod)
  3404   apply(auto simp add: lookupb_freshness forget fresh_atm fresh_prod)
  3397 done
  3405   done
  3398 
  3406 
  3399 lemma lookup_fresh:
  3407 lemma lookup_fresh:
  3400   assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
  3408   assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
  3401   shows "lookup x c \<theta>_n \<theta>_c = lookupa x c \<theta>_c"
  3409   shows "lookup x c \<theta>_n \<theta>_c = lookupa x c \<theta>_c"
  3402 using a
  3410   using a
  3403 apply(induct \<theta>_n)
  3411   apply(induct \<theta>_n)
  3404 apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3412    apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3405 done
  3413   done
  3406 
  3414 
  3407 lemma lookup_unicity:
  3415 lemma lookup_unicity:
  3408   assumes a: "lookup x a \<theta>_n \<theta>_c= Ax y b" "b\<sharp>(\<theta>_c,\<theta>_n)" "y\<sharp>(\<theta>_c,\<theta>_n)"
  3416   assumes a: "lookup x a \<theta>_n \<theta>_c= Ax y b" "b\<sharp>(\<theta>_c,\<theta>_n)" "y\<sharp>(\<theta>_c,\<theta>_n)"
  3409   shows "x=y \<and> a=b"
  3417   shows "x=y \<and> a=b"
  3410 using a
  3418   using a
  3411 apply(induct \<theta>_n)
  3419   apply(induct \<theta>_n)
  3412 apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
  3420    apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
  3413 apply(drule lookupa_unicity)
  3421      apply(drule lookupa_unicity)
  3414 apply(simp)+
  3422        apply(simp)+
  3415 apply(drule lookupa_unicity)
  3423     apply(drule lookupa_unicity)
  3416 apply(simp)+
  3424       apply(simp)+
  3417 apply(case_tac "x=aa")
  3425    apply(case_tac "x=aa")
  3418 apply(auto)
  3426     apply(auto)
  3419 apply(drule lookupb_unicity)
  3427    apply(drule lookupb_unicity)
  3420 apply(simp add: fresh_atm)
  3428      apply(simp add: fresh_atm)
  3421 apply(simp)
  3429     apply(simp)
  3422 apply(simp)
  3430    apply(simp)
  3423 apply(case_tac "x=aa")
  3431   apply(case_tac "x=aa")
  3424 apply(auto)
  3432    apply(auto)
  3425 apply(drule lookupb_unicity)
  3433   apply(drule lookupb_unicity)
  3426 apply(simp add: fresh_atm)
  3434     apply(simp add: fresh_atm)
  3427 apply(simp)
  3435    apply(simp)
  3428 apply(simp)
  3436   apply(simp)
  3429 done
  3437   done
  3430 
  3438 
  3431 lemma lookup_freshness:
  3439 lemma lookup_freshness:
  3432   fixes a::"coname"
  3440   fixes a::"coname"
  3433   and   x::"name"
  3441     and   x::"name"
  3434   shows "a\<sharp>(c,\<theta>_c,\<theta>_n) \<Longrightarrow> a\<sharp>lookup y c \<theta>_n \<theta>_c"
  3442   shows "a\<sharp>(c,\<theta>_c,\<theta>_n) \<Longrightarrow> a\<sharp>lookup y c \<theta>_n \<theta>_c"
  3435   and   "x\<sharp>(y,\<theta>_c,\<theta>_n) \<Longrightarrow> x\<sharp>lookup y c \<theta>_n \<theta>_c"   
  3443     and   "x\<sharp>(y,\<theta>_c,\<theta>_n) \<Longrightarrow> x\<sharp>lookup y c \<theta>_n \<theta>_c"   
  3436 apply(induct \<theta>_n)
  3444    apply(induct \<theta>_n)
  3437 apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3445      apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3438 apply(simp add: fresh_atm fresh_prod lookupa_freshness)
  3446      apply(simp add: fresh_atm fresh_prod lookupa_freshness)
  3439 apply(simp add: fresh_atm fresh_prod lookupa_freshness)
  3447     apply(simp add: fresh_atm fresh_prod lookupa_freshness)
  3440 apply(simp add: fresh_atm fresh_prod lookupb_freshness)
  3448    apply(simp add: fresh_atm fresh_prod lookupb_freshness)
  3441 apply(simp add: fresh_atm fresh_prod lookupb_freshness)
  3449   apply(simp add: fresh_atm fresh_prod lookupb_freshness)
  3442 done
  3450   done
  3443 
  3451 
  3444 lemma lookupc_freshness:
  3452 lemma lookupc_freshness:
  3445   fixes a::"coname"
  3453   fixes a::"coname"
  3446   and   x::"name"
  3454     and   x::"name"
  3447   shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupc y c \<theta>_c"
  3455   shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupc y c \<theta>_c"
  3448   and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupc y c \<theta>_c"
  3456     and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupc y c \<theta>_c"
  3449 apply(induct \<theta>_c)
  3457    apply(induct \<theta>_c)
  3450 apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3458      apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3451 apply(rule rename_fresh)
  3459    apply(rule rename_fresh)
  3452 apply(simp add: fresh_atm)
  3460    apply(simp add: fresh_atm)
  3453 apply(rule rename_fresh)
  3461   apply(rule rename_fresh)
  3454 apply(simp add: fresh_atm)
  3462   apply(simp add: fresh_atm)
  3455 done
  3463   done
  3456 
  3464 
  3457 lemma lookupc_fresh:
  3465 lemma lookupc_fresh:
  3458   assumes a: "y\<sharp>\<theta>_n"
  3466   assumes a: "y\<sharp>\<theta>_n"
  3459   shows "lookupc y a \<theta>_n = Ax y a"
  3467   shows "lookupc y a \<theta>_n = Ax y a"
  3460 using a
  3468   using a
  3461 apply(induct \<theta>_n)
  3469   apply(induct \<theta>_n)
  3462 apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3470    apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3463 done
  3471   done
  3464 
  3472 
  3465 lemma lookupc_nmaps:
  3473 lemma lookupc_nmaps:
  3466   assumes a: "\<theta>_n nmaps x to Some (c,P)"
  3474   assumes a: "\<theta>_n nmaps x to Some (c,P)"
  3467   shows "lookupc x a \<theta>_n = P[c\<turnstile>c>a]"
  3475   shows "lookupc x a \<theta>_n = P[c\<turnstile>c>a]"
  3468 using a
  3476   using a
  3469 apply(induct \<theta>_n)
  3477   apply(induct \<theta>_n)
  3470 apply(auto)
  3478    apply(auto)
  3471 done 
  3479   done 
  3472 
  3480 
  3473 lemma lookupc_unicity:
  3481 lemma lookupc_unicity:
  3474   assumes a: "lookupc y a \<theta>_n = Ax x b" "x\<sharp>\<theta>_n"
  3482   assumes a: "lookupc y a \<theta>_n = Ax x b" "x\<sharp>\<theta>_n"
  3475   shows "y=x"
  3483   shows "y=x"
  3476 using a
  3484   using a
  3477 apply(induct \<theta>_n)
  3485   apply(induct \<theta>_n)
  3478 apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
  3486    apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
  3479 apply(case_tac "y=aa")
  3487   apply(case_tac "y=aa")
  3480 apply(auto)
  3488    apply(auto)
  3481 apply(subgoal_tac "x\<sharp>(ba[aa\<turnstile>c>a])")
  3489   apply(subgoal_tac "x\<sharp>(ba[aa\<turnstile>c>a])")
  3482 apply(simp add: fresh_atm)
  3490    apply(simp add: fresh_atm)
  3483 apply(rule rename_fresh)
  3491   apply(rule rename_fresh)
  3484 apply(simp)
  3492   apply(simp)
  3485 done
  3493   done
  3486 
  3494 
  3487 lemma lookupd_fresh:
  3495 lemma lookupd_fresh:
  3488   assumes a: "a\<sharp>\<theta>_c"
  3496   assumes a: "a\<sharp>\<theta>_c"
  3489   shows "lookupd y a \<theta>_c = Ax y a"
  3497   shows "lookupd y a \<theta>_c = Ax y a"
  3490 using a
  3498   using a
  3491 apply(induct \<theta>_c)
  3499   apply(induct \<theta>_c)
  3492 apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3500    apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
  3493 done 
  3501   done 
  3494 
  3502 
  3495 lemma lookupd_unicity:
  3503 lemma lookupd_unicity:
  3496   assumes a: "lookupd y a \<theta>_c = Ax y b" "b\<sharp>\<theta>_c"
  3504   assumes a: "lookupd y a \<theta>_c = Ax y b" "b\<sharp>\<theta>_c"
  3497   shows "a=b"
  3505   shows "a=b"
  3498 using a
  3506   using a
  3499 apply(induct \<theta>_c)
  3507   apply(induct \<theta>_c)
  3500 apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
  3508    apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
  3501 apply(case_tac "a=aa")
  3509   apply(case_tac "a=aa")
  3502 apply(auto)
  3510    apply(auto)
  3503 apply(subgoal_tac "b\<sharp>(ba[aa\<turnstile>n>y])")
  3511   apply(subgoal_tac "b\<sharp>(ba[aa\<turnstile>n>y])")
  3504 apply(simp add: fresh_atm)
  3512    apply(simp add: fresh_atm)
  3505 apply(rule rename_fresh)
  3513   apply(rule rename_fresh)
  3506 apply(simp)
  3514   apply(simp)
  3507 done
  3515   done
  3508 
  3516 
  3509 lemma lookupd_freshness:
  3517 lemma lookupd_freshness:
  3510   fixes a::"coname"
  3518   fixes a::"coname"
  3511   and   x::"name"
  3519     and   x::"name"
  3512   shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupd y c \<theta>_c"
  3520   shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupd y c \<theta>_c"
  3513   and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupd y c \<theta>_c"
  3521     and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupd y c \<theta>_c"
  3514 apply(induct \<theta>_c)
  3522    apply(induct \<theta>_c)
  3515 apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3523      apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
  3516 apply(rule rename_fresh)
  3524    apply(rule rename_fresh)
  3517 apply(simp add: fresh_atm)
  3525    apply(simp add: fresh_atm)
  3518 apply(rule rename_fresh)
  3526   apply(rule rename_fresh)
  3519 apply(simp add: fresh_atm)
  3527   apply(simp add: fresh_atm)
  3520 done
  3528   done
  3521 
  3529 
  3522 lemma lookupd_cmaps:
  3530 lemma lookupd_cmaps:
  3523   assumes a: "\<theta>_c cmaps a to Some (x,P)"
  3531   assumes a: "\<theta>_c cmaps a to Some (x,P)"
  3524   shows "lookupd y a \<theta>_c = P[x\<turnstile>n>y]"
  3532   shows "lookupd y a \<theta>_c = P[x\<turnstile>n>y]"
  3525 using a
  3533   using a
  3526 apply(induct \<theta>_c)
  3534   apply(induct \<theta>_c)
  3527 apply(auto)
  3535    apply(auto)
  3528 done 
  3536   done 
  3529 
  3537 
  3530 nominal_primrec (freshness_context: "\<theta>_n::(name\<times>coname\<times>trm)")
  3538 nominal_primrec (freshness_context: "\<theta>_n::(name\<times>coname\<times>trm)")
  3531   stn :: "trm\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm" 
  3539   stn :: "trm\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm" 
  3532 where
  3540   where
  3533   "stn (Ax x a) \<theta>_n = lookupc x a \<theta>_n"
  3541     "stn (Ax x a) \<theta>_n = lookupc x a \<theta>_n"
  3534 | "\<lbrakk>a\<sharp>(N,\<theta>_n);x\<sharp>(M,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (Cut <a>.M (x).N) \<theta>_n = (Cut <a>.M (x).N)" 
  3542   | "\<lbrakk>a\<sharp>(N,\<theta>_n);x\<sharp>(M,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (Cut <a>.M (x).N) \<theta>_n = (Cut <a>.M (x).N)" 
  3535 | "x\<sharp>\<theta>_n \<Longrightarrow> stn (NotR (x).M a) \<theta>_n = (NotR (x).M a)"
  3543   | "x\<sharp>\<theta>_n \<Longrightarrow> stn (NotR (x).M a) \<theta>_n = (NotR (x).M a)"
  3536 | "a\<sharp>\<theta>_n \<Longrightarrow>stn (NotL <a>.M x) \<theta>_n = (NotL <a>.M x)"
  3544   | "a\<sharp>\<theta>_n \<Longrightarrow>stn (NotL <a>.M x) \<theta>_n = (NotL <a>.M x)"
  3537 | "\<lbrakk>a\<sharp>(N,d,b,\<theta>_n);b\<sharp>(M,d,a,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (AndR <a>.M <b>.N d) \<theta>_n = (AndR <a>.M <b>.N d)"
  3545   | "\<lbrakk>a\<sharp>(N,d,b,\<theta>_n);b\<sharp>(M,d,a,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (AndR <a>.M <b>.N d) \<theta>_n = (AndR <a>.M <b>.N d)"
  3538 | "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL1 (x).M z) \<theta>_n = (AndL1 (x).M z)"
  3546   | "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL1 (x).M z) \<theta>_n = (AndL1 (x).M z)"
  3539 | "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL2 (x).M z) \<theta>_n = (AndL2 (x).M z)"
  3547   | "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL2 (x).M z) \<theta>_n = (AndL2 (x).M z)"
  3540 | "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR1 <a>.M b) \<theta>_n = (OrR1 <a>.M b)"
  3548   | "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR1 <a>.M b) \<theta>_n = (OrR1 <a>.M b)"
  3541 | "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR2 <a>.M b) \<theta>_n = (OrR2 <a>.M b)"
  3549   | "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR2 <a>.M b) \<theta>_n = (OrR2 <a>.M b)"
  3542 | "\<lbrakk>x\<sharp>(N,z,u,\<theta>_n);u\<sharp>(M,z,x,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (OrL (x).M (u).N z) \<theta>_n = (OrL (x).M (u).N z)"
  3550   | "\<lbrakk>x\<sharp>(N,z,u,\<theta>_n);u\<sharp>(M,z,x,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (OrL (x).M (u).N z) \<theta>_n = (OrL (x).M (u).N z)"
  3543 | "\<lbrakk>a\<sharp>(b,\<theta>_n);x\<sharp>\<theta>_n\<rbrakk> \<Longrightarrow> stn (ImpR (x).<a>.M b) \<theta>_n = (ImpR (x).<a>.M b)"
  3551   | "\<lbrakk>a\<sharp>(b,\<theta>_n);x\<sharp>\<theta>_n\<rbrakk> \<Longrightarrow> stn (ImpR (x).<a>.M b) \<theta>_n = (ImpR (x).<a>.M b)"
  3544 | "\<lbrakk>a\<sharp>(N,\<theta>_n);x\<sharp>(M,z,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (ImpL <a>.M (x).N z) \<theta>_n = (ImpL <a>.M (x).N z)"
  3552   | "\<lbrakk>a\<sharp>(N,\<theta>_n);x\<sharp>(M,z,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (ImpL <a>.M (x).N z) \<theta>_n = (ImpL <a>.M (x).N z)"
  3545 apply(finite_guess)+
  3553                        apply(finite_guess)+
  3546 apply(rule TrueI)+
  3554                        apply(rule TrueI)+
  3547 apply(simp add: abs_fresh abs_supp fin_supp)+
  3555                        apply(simp add: abs_fresh abs_supp fin_supp)+
  3548 apply(fresh_guess)+
  3556                        apply(fresh_guess)+
  3549 done
  3557   done
  3550 
  3558 
  3551 nominal_primrec (freshness_context: "\<theta>_c::(coname\<times>name\<times>trm)")
  3559 nominal_primrec (freshness_context: "\<theta>_c::(coname\<times>name\<times>trm)")
  3552   stc :: "trm\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm" 
  3560   stc :: "trm\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm" 
  3553 where
  3561   where
  3554   "stc (Ax x a) \<theta>_c = lookupd x a \<theta>_c"
  3562     "stc (Ax x a) \<theta>_c = lookupd x a \<theta>_c"
  3555 | "\<lbrakk>a\<sharp>(N,\<theta>_c);x\<sharp>(M,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (Cut <a>.M (x).N) \<theta>_c = (Cut <a>.M (x).N)" 
  3563   | "\<lbrakk>a\<sharp>(N,\<theta>_c);x\<sharp>(M,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (Cut <a>.M (x).N) \<theta>_c = (Cut <a>.M (x).N)" 
  3556 | "x\<sharp>\<theta>_c \<Longrightarrow> stc (NotR (x).M a) \<theta>_c = (NotR (x).M a)"
  3564   | "x\<sharp>\<theta>_c \<Longrightarrow> stc (NotR (x).M a) \<theta>_c = (NotR (x).M a)"
  3557 | "a\<sharp>\<theta>_c \<Longrightarrow> stc (NotL <a>.M x) \<theta>_c = (NotL <a>.M x)"
  3565   | "a\<sharp>\<theta>_c \<Longrightarrow> stc (NotL <a>.M x) \<theta>_c = (NotL <a>.M x)"
  3558 | "\<lbrakk>a\<sharp>(N,d,b,\<theta>_c);b\<sharp>(M,d,a,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (AndR <a>.M <b>.N d) \<theta>_c = (AndR <a>.M <b>.N d)"
  3566   | "\<lbrakk>a\<sharp>(N,d,b,\<theta>_c);b\<sharp>(M,d,a,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (AndR <a>.M <b>.N d) \<theta>_c = (AndR <a>.M <b>.N d)"
  3559 | "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL1 (x).M z) \<theta>_c = (AndL1 (x).M z)"
  3567   | "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL1 (x).M z) \<theta>_c = (AndL1 (x).M z)"
  3560 | "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL2 (x).M z) \<theta>_c = (AndL2 (x).M z)"
  3568   | "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL2 (x).M z) \<theta>_c = (AndL2 (x).M z)"
  3561 | "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR1 <a>.M b) \<theta>_c = (OrR1 <a>.M b)"
  3569   | "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR1 <a>.M b) \<theta>_c = (OrR1 <a>.M b)"
  3562 | "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR2 <a>.M b) \<theta>_c = (OrR2 <a>.M b)"
  3570   | "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR2 <a>.M b) \<theta>_c = (OrR2 <a>.M b)"
  3563 | "\<lbrakk>x\<sharp>(N,z,u,\<theta>_c);u\<sharp>(M,z,x,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (OrL (x).M (u).N z) \<theta>_c = (OrL (x).M (u).N z)"
  3571   | "\<lbrakk>x\<sharp>(N,z,u,\<theta>_c);u\<sharp>(M,z,x,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (OrL (x).M (u).N z) \<theta>_c = (OrL (x).M (u).N z)"
  3564 | "\<lbrakk>a\<sharp>(b,\<theta>_c);x\<sharp>\<theta>_c\<rbrakk> \<Longrightarrow> stc (ImpR (x).<a>.M b) \<theta>_c = (ImpR (x).<a>.M b)"
  3572   | "\<lbrakk>a\<sharp>(b,\<theta>_c);x\<sharp>\<theta>_c\<rbrakk> \<Longrightarrow> stc (ImpR (x).<a>.M b) \<theta>_c = (ImpR (x).<a>.M b)"
  3565 | "\<lbrakk>a\<sharp>(N,\<theta>_c);x\<sharp>(M,z,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (ImpL <a>.M (x).N z) \<theta>_c = (ImpL <a>.M (x).N z)"
  3573   | "\<lbrakk>a\<sharp>(N,\<theta>_c);x\<sharp>(M,z,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (ImpL <a>.M (x).N z) \<theta>_c = (ImpL <a>.M (x).N z)"
  3566 apply(finite_guess)+
  3574                        apply(finite_guess)+
  3567 apply(rule TrueI)+
  3575                        apply(rule TrueI)+
  3568 apply(simp add: abs_fresh abs_supp fin_supp)+
  3576                        apply(simp add: abs_fresh abs_supp fin_supp)+
  3569 apply(fresh_guess)+
  3577                        apply(fresh_guess)+
  3570 done
  3578   done
  3571 
  3579 
  3572 lemma stn_eqvt[eqvt]:
  3580 lemma stn_eqvt[eqvt]:
  3573   fixes pi1::"name prm"
  3581   fixes pi1::"name prm"
  3574   and   pi2::"coname prm"
  3582     and   pi2::"coname prm"
  3575   shows "(pi1\<bullet>(stn M \<theta>_n)) = stn (pi1\<bullet>M) (pi1\<bullet>\<theta>_n)"
  3583   shows "(pi1\<bullet>(stn M \<theta>_n)) = stn (pi1\<bullet>M) (pi1\<bullet>\<theta>_n)"
  3576   and   "(pi2\<bullet>(stn M \<theta>_n)) = stn (pi2\<bullet>M) (pi2\<bullet>\<theta>_n)"
  3584     and   "(pi2\<bullet>(stn M \<theta>_n)) = stn (pi2\<bullet>M) (pi2\<bullet>\<theta>_n)"
  3577 apply -
  3585    apply -
  3578 apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
  3586    apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
  3579 apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3587               apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3580 apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
  3588   apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
  3581 apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3589              apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3582 done
  3590   done
  3583 
  3591 
  3584 lemma stc_eqvt[eqvt]:
  3592 lemma stc_eqvt[eqvt]:
  3585   fixes pi1::"name prm"
  3593   fixes pi1::"name prm"
  3586   and   pi2::"coname prm"
  3594     and   pi2::"coname prm"
  3587   shows "(pi1\<bullet>(stc M \<theta>_c)) = stc (pi1\<bullet>M) (pi1\<bullet>\<theta>_c)"
  3595   shows "(pi1\<bullet>(stc M \<theta>_c)) = stc (pi1\<bullet>M) (pi1\<bullet>\<theta>_c)"
  3588   and   "(pi2\<bullet>(stc M \<theta>_c)) = stc (pi2\<bullet>M) (pi2\<bullet>\<theta>_c)"
  3596     and   "(pi2\<bullet>(stc M \<theta>_c)) = stc (pi2\<bullet>M) (pi2\<bullet>\<theta>_c)"
  3589 apply -
  3597    apply -
  3590 apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
  3598    apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
  3591 apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3599               apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3592 apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
  3600   apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
  3593 apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3601              apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
  3594 done
  3602   done
  3595 
  3603 
  3596 lemma stn_fresh:
  3604 lemma stn_fresh:
  3597   fixes a::"coname"
  3605   fixes a::"coname"
  3598   and   x::"name"
  3606     and   x::"name"
  3599   shows "a\<sharp>(\<theta>_n,M) \<Longrightarrow> a\<sharp>stn M \<theta>_n"
  3607   shows "a\<sharp>(\<theta>_n,M) \<Longrightarrow> a\<sharp>stn M \<theta>_n"
  3600   and   "x\<sharp>(\<theta>_n,M) \<Longrightarrow> x\<sharp>stn M \<theta>_n"
  3608     and   "x\<sharp>(\<theta>_n,M) \<Longrightarrow> x\<sharp>stn M \<theta>_n"
  3601 apply(nominal_induct M avoiding: \<theta>_n a x rule: trm.strong_induct)
  3609    apply(nominal_induct M avoiding: \<theta>_n a x rule: trm.strong_induct)
  3602 apply(auto simp add: abs_fresh fresh_prod fresh_atm)
  3610                        apply(auto simp add: abs_fresh fresh_prod fresh_atm)
  3603 apply(rule lookupc_freshness)
  3611    apply(rule lookupc_freshness)
  3604 apply(simp add: fresh_atm)
  3612    apply(simp add: fresh_atm)
  3605 apply(rule lookupc_freshness)
  3613   apply(rule lookupc_freshness)
  3606 apply(simp add: fresh_atm)
  3614   apply(simp add: fresh_atm)
  3607 done
  3615   done
  3608 
  3616 
  3609 lemma stc_fresh:
  3617 lemma stc_fresh:
  3610   fixes a::"coname"
  3618   fixes a::"coname"
  3611   and   x::"name"
  3619     and   x::"name"
  3612   shows "a\<sharp>(\<theta>_c,M) \<Longrightarrow> a\<sharp>stc M \<theta>_c"
  3620   shows "a\<sharp>(\<theta>_c,M) \<Longrightarrow> a\<sharp>stc M \<theta>_c"
  3613   and   "x\<sharp>(\<theta>_c,M) \<Longrightarrow> x\<sharp>stc M \<theta>_c"
  3621     and   "x\<sharp>(\<theta>_c,M) \<Longrightarrow> x\<sharp>stc M \<theta>_c"
  3614 apply(nominal_induct M avoiding: \<theta>_c a x rule: trm.strong_induct)
  3622    apply(nominal_induct M avoiding: \<theta>_c a x rule: trm.strong_induct)
  3615 apply(auto simp add: abs_fresh fresh_prod fresh_atm)
  3623                        apply(auto simp add: abs_fresh fresh_prod fresh_atm)
  3616 apply(rule lookupd_freshness)
  3624    apply(rule lookupd_freshness)
  3617 apply(simp add: fresh_atm)
  3625    apply(simp add: fresh_atm)
  3618 apply(rule lookupd_freshness)
  3626   apply(rule lookupd_freshness)
  3619 apply(simp add: fresh_atm)
  3627   apply(simp add: fresh_atm)
  3620 done
  3628   done
  3621 
  3629 
  3622 lemma case_option_eqvt1[eqvt_force]:
  3630 lemma case_option_eqvt1[eqvt_force]:
  3623   fixes pi1::"name prm"
  3631   fixes pi1::"name prm"
  3624   and   pi2::"coname prm"
  3632     and   pi2::"coname prm"
  3625   and   B::"(name\<times>trm) option"
  3633     and   B::"(name\<times>trm) option"
  3626   and   r::"trm"
  3634     and   r::"trm"
  3627   shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
  3635   shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
  3628               (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
  3636               (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
  3629   and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
  3637     and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
  3630               (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
  3638               (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
  3631 apply(cases "B")
  3639    apply(cases "B")
  3632 apply(auto)
  3640     apply(auto)
  3633 apply(perm_simp)
  3641    apply(perm_simp)
  3634 apply(cases "B")
  3642   apply(cases "B")
  3635 apply(auto)
  3643    apply(auto)
  3636 apply(perm_simp)
  3644   apply(perm_simp)
  3637 done
  3645   done
  3638 
  3646 
  3639 lemma case_option_eqvt2[eqvt_force]:
  3647 lemma case_option_eqvt2[eqvt_force]:
  3640   fixes pi1::"name prm"
  3648   fixes pi1::"name prm"
  3641   and   pi2::"coname prm"
  3649     and   pi2::"coname prm"
  3642   and   B::"(coname\<times>trm) option"
  3650     and   B::"(coname\<times>trm) option"
  3643   and   r::"trm"
  3651     and   r::"trm"
  3644   shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
  3652   shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
  3645               (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
  3653               (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
  3646   and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
  3654     and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
  3647               (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
  3655               (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
  3648 apply(cases "B")
  3656    apply(cases "B")
  3649 apply(auto)
  3657     apply(auto)
  3650 apply(perm_simp)
  3658    apply(perm_simp)
  3651 apply(cases "B")
  3659   apply(cases "B")
  3652 apply(auto)
  3660    apply(auto)
  3653 apply(perm_simp)
  3661   apply(perm_simp)
  3654 done
  3662   done
  3655 
  3663 
  3656 nominal_primrec (freshness_context: "(\<theta>_n::(name\<times>coname\<times>trm) list,\<theta>_c::(coname\<times>name\<times>trm) list)")
  3664 nominal_primrec (freshness_context: "(\<theta>_n::(name\<times>coname\<times>trm) list,\<theta>_c::(coname\<times>name\<times>trm) list)")
  3657   psubst :: "(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm\<Rightarrow>trm" ("_,_<_>" [100,100,100] 100) 
  3665   psubst :: "(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm\<Rightarrow>trm" ("_,_<_>" [100,100,100] 100) 
  3658 where
  3666   where
  3659   "\<theta>_n,\<theta>_c<Ax x a> = lookup x a \<theta>_n \<theta>_c" 
  3667     "\<theta>_n,\<theta>_c<Ax x a> = lookup x a \<theta>_n \<theta>_c" 
  3660 | "\<lbrakk>a\<sharp>(N,\<theta>_n,\<theta>_c);x\<sharp>(M,\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> \<theta>_n,\<theta>_c<Cut <a>.M (x).N> = 
  3668   | "\<lbrakk>a\<sharp>(N,\<theta>_n,\<theta>_c);x\<sharp>(M,\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> \<theta>_n,\<theta>_c<Cut <a>.M (x).N> = 
  3661    Cut <a>.(if \<exists>x. M=Ax x a then stn M \<theta>_n else \<theta>_n,\<theta>_c<M>) 
  3669    Cut <a>.(if \<exists>x. M=Ax x a then stn M \<theta>_n else \<theta>_n,\<theta>_c<M>) 
  3662        (x).(if \<exists>a. N=Ax x a then stc N \<theta>_c else \<theta>_n,\<theta>_c<N>)" 
  3670        (x).(if \<exists>a. N=Ax x a then stc N \<theta>_c else \<theta>_n,\<theta>_c<N>)" 
  3663 | "x\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotR (x).M a> = 
  3671   | "x\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotR (x).M a> = 
  3664   (case (findc \<theta>_c a) of 
  3672   (case (findc \<theta>_c a) of 
  3665        Some (u,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).(\<theta>_n,\<theta>_c<M>) a' (u).P) 
  3673        Some (u,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).(\<theta>_n,\<theta>_c<M>) a' (u).P) 
  3666      | None \<Rightarrow> NotR (x).(\<theta>_n,\<theta>_c<M>) a)"
  3674      | None \<Rightarrow> NotR (x).(\<theta>_n,\<theta>_c<M>) a)"
  3667 | "a\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotL <a>.M x> = 
  3675   | "a\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotL <a>.M x> = 
  3668   (case (findn \<theta>_n x) of 
  3676   (case (findn \<theta>_n x) of 
  3669        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>x'. Cut <c>.P (x').(NotL <a>.(\<theta>_n,\<theta>_c<M>) x')) 
  3677        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>x'. Cut <c>.P (x').(NotL <a>.(\<theta>_n,\<theta>_c<M>) x')) 
  3670      | None \<Rightarrow> NotL <a>.(\<theta>_n,\<theta>_c<M>) x)"
  3678      | None \<Rightarrow> NotL <a>.(\<theta>_n,\<theta>_c<M>) x)"
  3671 | "\<lbrakk>a\<sharp>(N,c,\<theta>_n,\<theta>_c);b\<sharp>(M,c,\<theta>_n,\<theta>_c);b\<noteq>a\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<AndR <a>.M <b>.N c>) = 
  3679   | "\<lbrakk>a\<sharp>(N,c,\<theta>_n,\<theta>_c);b\<sharp>(M,c,\<theta>_n,\<theta>_c);b\<noteq>a\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<AndR <a>.M <b>.N c>) = 
  3672   (case (findc \<theta>_c c) of 
  3680   (case (findc \<theta>_c c) of 
  3673        Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.(\<theta>_n,\<theta>_c<M>) <b>.(\<theta>_n,\<theta>_c<N>) a') (x).P)
  3681        Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.(\<theta>_n,\<theta>_c<M>) <b>.(\<theta>_n,\<theta>_c<N>) a') (x).P)
  3674      | None \<Rightarrow> AndR <a>.(\<theta>_n,\<theta>_c<M>) <b>.(\<theta>_n,\<theta>_c<N>) c)"
  3682      | None \<Rightarrow> AndR <a>.(\<theta>_n,\<theta>_c<M>) <b>.(\<theta>_n,\<theta>_c<N>) c)"
  3675 | "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL1 (x).M z>) = 
  3683   | "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL1 (x).M z>) = 
  3676   (case (findn \<theta>_n z) of 
  3684   (case (findn \<theta>_n z) of 
  3677        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).(\<theta>_n,\<theta>_c<M>) z') 
  3685        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).(\<theta>_n,\<theta>_c<M>) z') 
  3678      | None \<Rightarrow> AndL1 (x).(\<theta>_n,\<theta>_c<M>) z)"
  3686      | None \<Rightarrow> AndL1 (x).(\<theta>_n,\<theta>_c<M>) z)"
  3679 | "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL2 (x).M z>) = 
  3687   | "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL2 (x).M z>) = 
  3680   (case (findn \<theta>_n z) of 
  3688   (case (findn \<theta>_n z) of 
  3681        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).(\<theta>_n,\<theta>_c<M>) z') 
  3689        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).(\<theta>_n,\<theta>_c<M>) z') 
  3682      | None \<Rightarrow> AndL2 (x).(\<theta>_n,\<theta>_c<M>) z)"
  3690      | None \<Rightarrow> AndL2 (x).(\<theta>_n,\<theta>_c<M>) z)"
  3683 | "\<lbrakk>x\<sharp>(N,z,\<theta>_n,\<theta>_c);u\<sharp>(M,z,\<theta>_n,\<theta>_c);x\<noteq>u\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<OrL (x).M (u).N z>) =
  3691   | "\<lbrakk>x\<sharp>(N,z,\<theta>_n,\<theta>_c);u\<sharp>(M,z,\<theta>_n,\<theta>_c);x\<noteq>u\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<OrL (x).M (u).N z>) =
  3684   (case (findn \<theta>_n z) of  
  3692   (case (findn \<theta>_n z) of  
  3685        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).(\<theta>_n,\<theta>_c<M>) (u).(\<theta>_n,\<theta>_c<N>) z') 
  3693        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).(\<theta>_n,\<theta>_c<M>) (u).(\<theta>_n,\<theta>_c<N>) z') 
  3686      | None \<Rightarrow> OrL (x).(\<theta>_n,\<theta>_c<M>) (u).(\<theta>_n,\<theta>_c<N>) z)"
  3694      | None \<Rightarrow> OrL (x).(\<theta>_n,\<theta>_c<M>) (u).(\<theta>_n,\<theta>_c<N>) z)"
  3687 | "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR1 <a>.M b>) = 
  3695   | "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR1 <a>.M b>) = 
  3688   (case (findc \<theta>_c b) of
  3696   (case (findc \<theta>_c b) of
  3689        Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
  3697        Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
  3690      | None \<Rightarrow> OrR1 <a>.(\<theta>_n,\<theta>_c<M>) b)"
  3698      | None \<Rightarrow> OrR1 <a>.(\<theta>_n,\<theta>_c<M>) b)"
  3691 | "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR2 <a>.M b>) = 
  3699   | "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR2 <a>.M b>) = 
  3692   (case (findc \<theta>_c b) of
  3700   (case (findc \<theta>_c b) of
  3693        Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
  3701        Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
  3694      | None \<Rightarrow> OrR2 <a>.(\<theta>_n,\<theta>_c<M>) b)"
  3702      | None \<Rightarrow> OrR2 <a>.(\<theta>_n,\<theta>_c<M>) b)"
  3695 | "\<lbrakk>a\<sharp>(b,\<theta>_n,\<theta>_c); x\<sharp>(\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<ImpR (x).<a>.M b>) = 
  3703   | "\<lbrakk>a\<sharp>(b,\<theta>_n,\<theta>_c); x\<sharp>(\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<ImpR (x).<a>.M b>) = 
  3696   (case (findc \<theta>_c b) of
  3704   (case (findc \<theta>_c b) of
  3697        Some (z,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) a' (z).P)
  3705        Some (z,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) a' (z).P)
  3698      | None \<Rightarrow> ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) b)"
  3706      | None \<Rightarrow> ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) b)"
  3699 | "\<lbrakk>a\<sharp>(N,\<theta>_n,\<theta>_c); x\<sharp>(z,M,\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<ImpL <a>.M (x).N z>) = 
  3707   | "\<lbrakk>a\<sharp>(N,\<theta>_n,\<theta>_c); x\<sharp>(z,M,\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<ImpL <a>.M (x).N z>) = 
  3700   (case (findn \<theta>_n z) of
  3708   (case (findn \<theta>_n z) of
  3701        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>) z') 
  3709        Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>) z') 
  3702      | None \<Rightarrow> ImpL <a>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>) z)"
  3710      | None \<Rightarrow> ImpL <a>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>) z)"
  3703 apply(finite_guess)+
  3711                        apply(finite_guess)+
  3704 apply(rule TrueI)+
  3712                        apply(rule TrueI)+
  3705 apply(simp add: abs_fresh stc_fresh)
  3713                        apply(simp add: abs_fresh stc_fresh)
  3706 apply(simp add: abs_fresh stn_fresh)
  3714                        apply(simp add: abs_fresh stn_fresh)
  3707 apply(case_tac "findc \<theta>_c x3")
  3715                        apply(case_tac "findc \<theta>_c x3")
  3708 apply(simp add: abs_fresh)
  3716                        apply(simp add: abs_fresh)
  3709 apply(auto)[1]
  3717                        apply(auto)[1]
  3710 apply(generate_fresh "coname")
  3718                        apply(generate_fresh "coname")
  3711 apply(fresh_fun_simp (no_asm))
  3719                        apply(fresh_fun_simp (no_asm))
  3712 apply(drule cmaps_fresh)
  3720                        apply(drule cmaps_fresh)
  3713 apply(auto simp add: fresh_prod)[1]
  3721                        apply(auto simp add: fresh_prod)[1]
  3714 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3722                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3715 apply(case_tac "findn \<theta>_n x3")
  3723                        apply(case_tac "findn \<theta>_n x3")
  3716 apply(simp add: abs_fresh)
  3724                        apply(simp add: abs_fresh)
  3717 apply(auto)[1]
  3725                        apply(auto)[1]
  3718 apply(generate_fresh "name")
  3726                        apply(generate_fresh "name")
  3719 apply(fresh_fun_simp (no_asm))
  3727                        apply(fresh_fun_simp (no_asm))
  3720 apply(drule nmaps_fresh)
  3728                        apply(drule nmaps_fresh)
  3721 apply(auto simp add: fresh_prod)[1]
  3729                        apply(auto simp add: fresh_prod)[1]
  3722 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3730                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3723 apply(case_tac "findc \<theta>_c x5")
  3731                        apply(case_tac "findc \<theta>_c x5")
  3724 apply(simp add: abs_fresh)
  3732                        apply(simp add: abs_fresh)
  3725 apply(auto)[1]
  3733                        apply(auto)[1]
  3726 apply(generate_fresh "coname")
  3734                        apply(generate_fresh "coname")
  3727 apply(fresh_fun_simp (no_asm))
  3735                        apply(fresh_fun_simp (no_asm))
  3728 apply(drule cmaps_fresh)
  3736                        apply(drule cmaps_fresh)
  3729 apply(auto simp add: fresh_prod)[1]
  3737                        apply(auto simp add: fresh_prod)[1]
  3730 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3738                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3731 apply(case_tac "findc \<theta>_c x5")
  3739                        apply(case_tac "findc \<theta>_c x5")
  3732 apply(simp add: abs_fresh)
  3740                        apply(simp add: abs_fresh)
  3733 apply(auto)[1]
  3741                        apply(auto)[1]
  3734 apply(generate_fresh "coname")
  3742                        apply(generate_fresh "coname")
  3735 apply(fresh_fun_simp (no_asm))
  3743                        apply(fresh_fun_simp (no_asm))
  3736 apply(drule_tac x="x3" in cmaps_fresh)
  3744                        apply(drule_tac x="x3" in cmaps_fresh)
  3737 apply(auto simp add: fresh_prod)[1]
  3745                        apply(auto simp add: fresh_prod)[1]
  3738 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3746                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3739 apply(case_tac "findn \<theta>_n x3")
  3747                        apply(case_tac "findn \<theta>_n x3")
  3740 apply(simp add: abs_fresh)
  3748                        apply(simp add: abs_fresh)
  3741 apply(auto)[1]
  3749                        apply(auto)[1]
  3742 apply(generate_fresh "name")
  3750                        apply(generate_fresh "name")
  3743 apply(fresh_fun_simp (no_asm))
  3751                        apply(fresh_fun_simp (no_asm))
  3744 apply(drule nmaps_fresh)
  3752                        apply(drule nmaps_fresh)
  3745 apply(auto simp add: fresh_prod)[1]
  3753                        apply(auto simp add: fresh_prod)[1]
  3746 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3754                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3747 apply(case_tac "findn \<theta>_n x3")
  3755                        apply(case_tac "findn \<theta>_n x3")
  3748 apply(simp add: abs_fresh)
  3756                        apply(simp add: abs_fresh)
  3749 apply(auto)[1]
  3757                        apply(auto)[1]
  3750 apply(generate_fresh "name")
  3758                        apply(generate_fresh "name")
  3751 apply(fresh_fun_simp (no_asm))
  3759                        apply(fresh_fun_simp (no_asm))
  3752 apply(drule nmaps_fresh)
  3760                        apply(drule nmaps_fresh)
  3753 apply(auto simp add: fresh_prod)[1]
  3761                        apply(auto simp add: fresh_prod)[1]
  3754 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3762                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3755 apply(case_tac "findc \<theta>_c x3")
  3763                        apply(case_tac "findc \<theta>_c x3")
  3756 apply(simp add: abs_fresh)
  3764                        apply(simp add: abs_fresh)
  3757 apply(auto)[1]
  3765                        apply(auto)[1]
  3758 apply(generate_fresh "coname")
  3766                        apply(generate_fresh "coname")
  3759 apply(fresh_fun_simp (no_asm))
  3767                        apply(fresh_fun_simp (no_asm))
  3760 apply(drule cmaps_fresh)
  3768                        apply(drule cmaps_fresh)
  3761 apply(auto simp add: fresh_prod)[1]
  3769                        apply(auto simp add: fresh_prod)[1]
  3762 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3770                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3763 apply(case_tac "findc \<theta>_c x3")
  3771                        apply(case_tac "findc \<theta>_c x3")
  3764 apply(simp add: abs_fresh)
  3772                        apply(simp add: abs_fresh)
  3765 apply(auto)[1]
  3773                        apply(auto)[1]
  3766 apply(generate_fresh "coname")
  3774                        apply(generate_fresh "coname")
  3767 apply(fresh_fun_simp (no_asm))
  3775                        apply(fresh_fun_simp (no_asm))
  3768 apply(drule cmaps_fresh)
  3776                        apply(drule cmaps_fresh)
  3769 apply(auto simp add: fresh_prod)[1]
  3777                        apply(auto simp add: fresh_prod)[1]
  3770 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3778                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3771 apply(case_tac "findn \<theta>_n x5")
  3779                        apply(case_tac "findn \<theta>_n x5")
  3772 apply(simp add: abs_fresh)
  3780                        apply(simp add: abs_fresh)
  3773 apply(auto)[1]
  3781                        apply(auto)[1]
  3774 apply(generate_fresh "name")
  3782                        apply(generate_fresh "name")
  3775 apply(fresh_fun_simp (no_asm))
  3783                        apply(fresh_fun_simp (no_asm))
  3776 apply(drule nmaps_fresh)
  3784                        apply(drule nmaps_fresh)
  3777 apply(auto simp add: fresh_prod)[1]
  3785                        apply(auto simp add: fresh_prod)[1]
  3778 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3786                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3779 apply(case_tac "findn \<theta>_n x5")
  3787                        apply(case_tac "findn \<theta>_n x5")
  3780 apply(simp add: abs_fresh)
  3788                        apply(simp add: abs_fresh)
  3781 apply(auto)[1]
  3789                        apply(auto)[1]
  3782 apply(generate_fresh "name")
  3790                        apply(generate_fresh "name")
  3783 apply(fresh_fun_simp (no_asm))
  3791                        apply(fresh_fun_simp (no_asm))
  3784 apply(drule_tac a="x3" in nmaps_fresh)
  3792                        apply(drule_tac a="x3" in nmaps_fresh)
  3785 apply(auto simp add: fresh_prod)[1]
  3793                        apply(auto simp add: fresh_prod)[1]
  3786 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3794                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3787 apply(case_tac "findc \<theta>_c x4")
  3795                        apply(case_tac "findc \<theta>_c x4")
  3788 apply(simp add: abs_fresh abs_supp fin_supp)
  3796                        apply(simp add: abs_fresh abs_supp fin_supp)
  3789 apply(auto)[1]
  3797                        apply(auto)[1]
  3790 apply(generate_fresh "coname")
  3798                        apply(generate_fresh "coname")
  3791 apply(fresh_fun_simp (no_asm))
  3799                        apply(fresh_fun_simp (no_asm))
  3792 apply(drule cmaps_fresh)
  3800                        apply(drule cmaps_fresh)
  3793 apply(auto simp add: fresh_prod)[1]
  3801                        apply(auto simp add: fresh_prod)[1]
  3794 apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
  3802                        apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
  3795 apply(case_tac "findc \<theta>_c x4")
  3803                        apply(case_tac "findc \<theta>_c x4")
  3796 apply(simp add: abs_fresh abs_supp fin_supp)
  3804                        apply(simp add: abs_fresh abs_supp fin_supp)
  3797 apply(auto)[1]
  3805                        apply(auto)[1]
  3798 apply(generate_fresh "coname")
  3806                        apply(generate_fresh "coname")
  3799 apply(fresh_fun_simp (no_asm))
  3807                        apply(fresh_fun_simp (no_asm))
  3800 apply(drule_tac x="x2" in cmaps_fresh)
  3808                        apply(drule_tac x="x2" in cmaps_fresh)
  3801 apply(auto simp add: fresh_prod)[1]
  3809                        apply(auto simp add: fresh_prod)[1]
  3802 apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
  3810                        apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
  3803 apply(case_tac "findn \<theta>_n x5")
  3811                        apply(case_tac "findn \<theta>_n x5")
  3804 apply(simp add: abs_fresh)
  3812                        apply(simp add: abs_fresh)
  3805 apply(auto)[1]
  3813                        apply(auto)[1]
  3806 apply(generate_fresh "name")
  3814                        apply(generate_fresh "name")
  3807 apply(fresh_fun_simp (no_asm))
  3815                        apply(fresh_fun_simp (no_asm))
  3808 apply(drule nmaps_fresh)
  3816                        apply(drule nmaps_fresh)
  3809 apply(auto simp add: fresh_prod)[1]
  3817                        apply(auto simp add: fresh_prod)[1]
  3810 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3818                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3811 apply(case_tac "findn \<theta>_n x5")
  3819                        apply(case_tac "findn \<theta>_n x5")
  3812 apply(simp add: abs_fresh)
  3820                        apply(simp add: abs_fresh)
  3813 apply(auto)[1]
  3821                        apply(auto)[1]
  3814 apply(generate_fresh "name")
  3822                        apply(generate_fresh "name")
  3815 apply(fresh_fun_simp (no_asm))
  3823                        apply(fresh_fun_simp (no_asm))
  3816 apply(drule_tac a="x3" in nmaps_fresh)
  3824                        apply(drule_tac a="x3" in nmaps_fresh)
  3817 apply(auto simp add: fresh_prod)[1]
  3825                        apply(auto simp add: fresh_prod)[1]
  3818 apply(simp add: abs_fresh fresh_prod fresh_atm)
  3826                        apply(simp add: abs_fresh fresh_prod fresh_atm)
  3819 apply(fresh_guess)+
  3827                        apply(fresh_guess)+
  3820 done
  3828   done
  3821 
  3829 
  3822 lemma case_cong:
  3830 lemma case_cong:
  3823   assumes a: "B1=B2" "x1=x2" "y1=y2"
  3831   assumes a: "B1=B2" "x1=x2" "y1=y2"
  3824   shows "(case B1 of None \<Rightarrow> x1 | Some (x,P) \<Rightarrow> y1 x P) = (case B2 of None \<Rightarrow> x2 | Some (x,P) \<Rightarrow> y2 x P)"
  3832   shows "(case B1 of None \<Rightarrow> x1 | Some (x,P) \<Rightarrow> y1 x P) = (case B2 of None \<Rightarrow> x2 | Some (x,P) \<Rightarrow> y2 x P)"
  3825 using a
  3833   using a
  3826 apply(auto)
  3834   apply(auto)
  3827 done
  3835   done
  3828 
  3836 
  3829 lemma find_maps:
  3837 lemma find_maps:
  3830   shows "\<theta>_c cmaps a to (findc \<theta>_c a)"
  3838   shows "\<theta>_c cmaps a to (findc \<theta>_c a)"
  3831   and   "\<theta>_n nmaps x to (findn \<theta>_n x)"
  3839     and   "\<theta>_n nmaps x to (findn \<theta>_n x)"
  3832 apply(auto)
  3840    apply(auto)
  3833 done
  3841   done
  3834 
  3842 
  3835 lemma psubst_eqvt[eqvt]:
  3843 lemma psubst_eqvt[eqvt]:
  3836   fixes pi1::"name prm"
  3844   fixes pi1::"name prm"
  3837   and   pi2::"coname prm"
  3845     and   pi2::"coname prm"
  3838   shows "pi1\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi1\<bullet>\<theta>_n),(pi1\<bullet>\<theta>_c)<(pi1\<bullet>M)>"
  3846   shows "pi1\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi1\<bullet>\<theta>_n),(pi1\<bullet>\<theta>_c)<(pi1\<bullet>M)>"
  3839   and   "pi2\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi2\<bullet>\<theta>_n),(pi2\<bullet>\<theta>_c)<(pi2\<bullet>M)>"
  3847     and   "pi2\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi2\<bullet>\<theta>_n),(pi2\<bullet>\<theta>_c)<(pi2\<bullet>M)>"
  3840 apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
  3848    apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
  3841 apply(auto simp add: eq_bij fresh_bij eqvts perm_pi_simp)
  3849                        apply(auto simp add: eq_bij fresh_bij eqvts perm_pi_simp)
  3842 apply(rule case_cong)
  3850                      apply(rule case_cong)
  3843 apply(rule find_maps)
  3851                        apply(rule find_maps)
  3844 apply(simp)
  3852                       apply(simp)
  3845 apply(perm_simp add: eqvts)
  3853                      apply(perm_simp add: eqvts)
  3846 apply(rule case_cong)
  3854                     apply(rule case_cong)
  3847 apply(rule find_maps)
  3855                       apply(rule find_maps)
  3848 apply(simp)
  3856                      apply(simp)
  3849 apply(perm_simp add: eqvts)
  3857                     apply(perm_simp add: eqvts)
  3850 apply(rule case_cong)
  3858                    apply(rule case_cong)
  3851 apply(rule find_maps)
  3859                      apply(rule find_maps)
  3852 apply(simp)
  3860                     apply(simp)
  3853 apply(perm_simp add: eqvts)
  3861                    apply(perm_simp add: eqvts)
  3854 apply(rule case_cong)
  3862                   apply(rule case_cong)
  3855 apply(rule find_maps)
  3863                     apply(rule find_maps)
  3856 apply(simp)
  3864                    apply(simp)
  3857 apply(perm_simp add: eqvts)
  3865                   apply(perm_simp add: eqvts)
  3858 apply(rule case_cong)
  3866                  apply(rule case_cong)
  3859 apply(rule find_maps)
  3867                    apply(rule find_maps)
  3860 apply(simp)
  3868                   apply(simp)
  3861 apply(perm_simp add: eqvts)
  3869                  apply(perm_simp add: eqvts)
  3862 apply(rule case_cong)
  3870                 apply(rule case_cong)
  3863 apply(rule find_maps)
  3871                   apply(rule find_maps)
  3864 apply(simp)
  3872                  apply(simp)
  3865 apply(perm_simp add: eqvts)
  3873                 apply(perm_simp add: eqvts)
  3866 apply(rule case_cong)
  3874                apply(rule case_cong)
  3867 apply(rule find_maps)
  3875                  apply(rule find_maps)
  3868 apply(simp)
  3876                 apply(simp)
  3869 apply(perm_simp add: eqvts)
  3877                apply(perm_simp add: eqvts)
  3870 apply(rule case_cong)
  3878               apply(rule case_cong)
  3871 apply(rule find_maps)
  3879                 apply(rule find_maps)
  3872 apply(simp)
  3880                apply(simp)
  3873 apply(perm_simp add: eqvts)
  3881               apply(perm_simp add: eqvts)
  3874 apply(rule case_cong)
  3882              apply(rule case_cong)
  3875 apply(rule find_maps)
  3883                apply(rule find_maps)
  3876 apply(simp)
  3884               apply(simp)
  3877 apply(perm_simp add: eqvts)
  3885              apply(perm_simp add: eqvts)
  3878 apply(rule case_cong)
  3886             apply(rule case_cong)
  3879 apply(rule find_maps)
  3887               apply(rule find_maps)
  3880 apply(simp)
  3888              apply(simp)
  3881 apply(perm_simp add: eqvts)
  3889             apply(perm_simp add: eqvts)
  3882 apply(rule case_cong)
  3890            apply(rule case_cong)
  3883 apply(rule find_maps)
  3891              apply(rule find_maps)
  3884 apply(simp)
  3892             apply(simp)
  3885 apply(perm_simp add: eqvts)
  3893            apply(perm_simp add: eqvts)
  3886 apply(rule case_cong)
  3894           apply(rule case_cong)
  3887 apply(rule find_maps)
  3895             apply(rule find_maps)
  3888 apply(simp)
  3896            apply(simp)
  3889 apply(perm_simp add: eqvts)
  3897           apply(perm_simp add: eqvts)
  3890 apply(rule case_cong)
  3898          apply(rule case_cong)
  3891 apply(rule find_maps)
  3899            apply(rule find_maps)
  3892 apply(simp)
  3900           apply(simp)
  3893 apply(perm_simp add: eqvts)
  3901          apply(perm_simp add: eqvts)
  3894 apply(rule case_cong)
  3902         apply(rule case_cong)
  3895 apply(rule find_maps)
  3903           apply(rule find_maps)
  3896 apply(simp)
  3904          apply(simp)
  3897 apply(perm_simp add: eqvts)
  3905         apply(perm_simp add: eqvts)
  3898 apply(rule case_cong)
  3906        apply(rule case_cong)
  3899 apply(rule find_maps)
  3907          apply(rule find_maps)
  3900 apply(simp)
  3908         apply(simp)
  3901 apply(perm_simp add: eqvts)
  3909        apply(perm_simp add: eqvts)
  3902 apply(rule case_cong)
  3910       apply(rule case_cong)
  3903 apply(rule find_maps)
  3911         apply(rule find_maps)
  3904 apply(simp)
  3912        apply(simp)
  3905 apply(perm_simp add: eqvts)
  3913       apply(perm_simp add: eqvts)
  3906 apply(rule case_cong)
  3914      apply(rule case_cong)
  3907 apply(rule find_maps)
  3915        apply(rule find_maps)
  3908 apply(simp)
  3916       apply(simp)
  3909 apply(perm_simp add: eqvts)
  3917      apply(perm_simp add: eqvts)
  3910 apply(rule case_cong)
  3918     apply(rule case_cong)
  3911 apply(rule find_maps)
  3919       apply(rule find_maps)
  3912 apply(simp)
  3920      apply(simp)
  3913 apply(perm_simp add: eqvts)
  3921     apply(perm_simp add: eqvts)
  3914 apply(rule case_cong)
  3922    apply(rule case_cong)
  3915 apply(rule find_maps)
  3923      apply(rule find_maps)
  3916 apply(simp)
  3924     apply(simp)
  3917 apply(perm_simp add: eqvts)
  3925    apply(perm_simp add: eqvts)
  3918 apply(rule case_cong)
  3926   apply(rule case_cong)
  3919 apply(rule find_maps)
  3927     apply(rule find_maps)
  3920 apply(simp)
  3928    apply(simp)
  3921 apply(perm_simp add: eqvts)
  3929   apply(perm_simp add: eqvts)
  3922 done
  3930   done
  3923 
  3931 
  3924 lemma ax_psubst:
  3932 lemma ax_psubst:
  3925   assumes a: "\<theta>_n,\<theta>_c<M> = Ax x a"
  3933   assumes a: "\<theta>_n,\<theta>_c<M> = Ax x a"
  3926   and     b: "a\<sharp>(\<theta>_n,\<theta>_c)" "x\<sharp>(\<theta>_n,\<theta>_c)"
  3934     and     b: "a\<sharp>(\<theta>_n,\<theta>_c)" "x\<sharp>(\<theta>_n,\<theta>_c)"
  3927   shows "M = Ax x a"
  3935   shows "M = Ax x a"
  3928 using a b
  3936   using a b
  3929 apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
  3937   apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
  3930 apply(auto)
  3938              apply(auto)
  3931 apply(drule lookup_unicity)
  3939             apply(drule lookup_unicity)
  3932 apply(simp)+
  3940               apply(simp)+
  3933 apply(case_tac "findc \<theta>_c coname")
  3941            apply(case_tac "findc \<theta>_c coname")
  3934 apply(simp)
  3942             apply(simp)
  3935 apply(auto)[1]
  3943            apply(auto)[1]
  3936 apply(generate_fresh "coname")
  3944            apply(generate_fresh "coname")
  3937 apply(fresh_fun_simp)
  3945            apply(fresh_fun_simp)
  3938 apply(simp)
  3946            apply(simp)
  3939 apply(case_tac "findn \<theta>_n name")
  3947           apply(case_tac "findn \<theta>_n name")
  3940 apply(simp)
  3948            apply(simp)
  3941 apply(auto)[1]
  3949           apply(auto)[1]
  3942 apply(generate_fresh "name")
  3950           apply(generate_fresh "name")
  3943 apply(fresh_fun_simp)
  3951           apply(fresh_fun_simp)
  3944 apply(simp)
  3952           apply(simp)
  3945 apply(case_tac "findc \<theta>_c coname3")
  3953          apply(case_tac "findc \<theta>_c coname3")
  3946 apply(simp)
  3954           apply(simp)
  3947 apply(auto)[1]
  3955          apply(auto)[1]
  3948 apply(generate_fresh "coname")
  3956          apply(generate_fresh "coname")
  3949 apply(fresh_fun_simp)
  3957          apply(fresh_fun_simp)
  3950 apply(simp)
  3958          apply(simp)
  3951 apply(case_tac "findn \<theta>_n name2")
  3959         apply(case_tac "findn \<theta>_n name2")
  3952 apply(simp)
  3960          apply(simp)
  3953 apply(auto)[1]
  3961         apply(auto)[1]
  3954 apply(generate_fresh "name")
  3962         apply(generate_fresh "name")
  3955 apply(fresh_fun_simp)
  3963         apply(fresh_fun_simp)
  3956 apply(simp)
  3964         apply(simp)
  3957 apply(case_tac "findn \<theta>_n name2")
  3965        apply(case_tac "findn \<theta>_n name2")
  3958 apply(simp)
  3966         apply(simp)
  3959 apply(auto)[1]
  3967        apply(auto)[1]
  3960 apply(generate_fresh "name")
  3968        apply(generate_fresh "name")
  3961 apply(fresh_fun_simp)
  3969        apply(fresh_fun_simp)
  3962 apply(simp)
  3970        apply(simp)
  3963 apply(case_tac "findc \<theta>_c coname2")
  3971       apply(case_tac "findc \<theta>_c coname2")
  3964 apply(simp)
  3972        apply(simp)
  3965 apply(auto)[1]
  3973       apply(auto)[1]
  3966 apply(generate_fresh "coname")
  3974       apply(generate_fresh "coname")
  3967 apply(fresh_fun_simp)
  3975       apply(fresh_fun_simp)
  3968 apply(simp)
  3976       apply(simp)
  3969 apply(case_tac "findc \<theta>_c coname2")
  3977      apply(case_tac "findc \<theta>_c coname2")
  3970 apply(simp)
  3978       apply(simp)
  3971 apply(auto)[1]
  3979      apply(auto)[1]
  3972 apply(generate_fresh "coname")
  3980      apply(generate_fresh "coname")
  3973 apply(fresh_fun_simp)
  3981      apply(fresh_fun_simp)
  3974 apply(simp)
  3982      apply(simp)
  3975 apply(case_tac "findn \<theta>_n name3")
  3983     apply(case_tac "findn \<theta>_n name3")
  3976 apply(simp)
  3984      apply(simp)
  3977 apply(auto)[1]
  3985     apply(auto)[1]
  3978 apply(generate_fresh "name")
  3986     apply(generate_fresh "name")
  3979 apply(fresh_fun_simp)
  3987     apply(fresh_fun_simp)
  3980 apply(simp)
  3988     apply(simp)
  3981 apply(case_tac "findc \<theta>_c coname2")
  3989    apply(case_tac "findc \<theta>_c coname2")
  3982 apply(simp)
  3990     apply(simp)
  3983 apply(auto)[1]
  3991    apply(auto)[1]
  3984 apply(generate_fresh "coname")
  3992    apply(generate_fresh "coname")
  3985 apply(fresh_fun_simp)
  3993    apply(fresh_fun_simp)
  3986 apply(simp)
  3994    apply(simp)
  3987 apply(case_tac "findn \<theta>_n name2")
  3995   apply(case_tac "findn \<theta>_n name2")
  3988 apply(simp)
  3996    apply(simp)
  3989 apply(auto)[1]
  3997   apply(auto)[1]
  3990 apply(generate_fresh "name")
  3998   apply(generate_fresh "name")
  3991 apply(fresh_fun_simp)
  3999   apply(fresh_fun_simp)
  3992 apply(simp)
  4000   apply(simp)
  3993 done
  4001   done
  3994 
  4002 
  3995 lemma better_Cut_substc1:
  4003 lemma better_Cut_substc1:
  3996   assumes a: "a\<sharp>(P,b)" "b\<sharp>N" 
  4004   assumes a: "a\<sharp>(P,b)" "b\<sharp>N" 
  3997   shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.(M{b:=(y).P}) (x).N"
  4005   shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.(M{b:=(y).P}) (x).N"
  3998 using a
  4006   using a
  3999 apply -
  4007   apply -
  4000 apply(generate_fresh "coname")
  4008   apply(generate_fresh "coname")
  4001 apply(generate_fresh "name")
  4009   apply(generate_fresh "name")
  4002 apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4010   apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4003 apply(simp)
  4011    apply(simp)
  4004 apply(rule trans)
  4012    apply(rule trans)
  4005 apply(rule better_Cut_substc)
  4013     apply(rule better_Cut_substc)
  4006 apply(simp)
  4014      apply(simp)
  4007 apply(simp add: abs_fresh)
  4015     apply(simp add: abs_fresh)
  4008 apply(auto)[1]
  4016    apply(auto)[1]
  4009 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
  4017     apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
  4010 apply(simp add: calc_atm fresh_atm)
  4018     apply(simp add: calc_atm fresh_atm)
  4011 apply(subgoal_tac"b\<sharp>([(ca,x)]\<bullet>N)")
  4019    apply(subgoal_tac"b\<sharp>([(ca,x)]\<bullet>N)")
  4012 apply(simp add: forget)
  4020     apply(simp add: forget)
  4013 apply(simp add: trm.inject)
  4021     apply(simp add: trm.inject)
  4014 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4022     apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4015 apply(perm_simp)
  4023     apply(perm_simp)
  4016 apply(simp add: fresh_left calc_atm)
  4024    apply(simp add: fresh_left calc_atm)
  4017 apply(simp add: trm.inject)
  4025   apply(simp add: trm.inject)
  4018 apply(rule conjI)
  4026   apply(rule conjI)
  4019 apply(rule sym)
  4027    apply(rule sym)
  4020 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4028    apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4021 apply(rule sym)
  4029   apply(rule sym)
  4022 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4030   apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4023 done
  4031   done
  4024 
  4032 
  4025 lemma better_Cut_substc2:
  4033 lemma better_Cut_substc2:
  4026   assumes a: "x\<sharp>(y,P)" "b\<sharp>(a,M)" "N\<noteq>Ax x b"
  4034   assumes a: "x\<sharp>(y,P)" "b\<sharp>(a,M)" "N\<noteq>Ax x b"
  4027   shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.M (x).(N{b:=(y).P})"
  4035   shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.M (x).(N{b:=(y).P})"
  4028 using a
  4036   using a
  4029 apply -
  4037   apply -
  4030 apply(generate_fresh "coname")
  4038   apply(generate_fresh "coname")
  4031 apply(generate_fresh "name")
  4039   apply(generate_fresh "name")
  4032 apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4040   apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4033 apply(simp)
  4041    apply(simp)
  4034 apply(rule trans)
  4042    apply(rule trans)
  4035 apply(rule better_Cut_substc)
  4043     apply(rule better_Cut_substc)
  4036 apply(simp)
  4044      apply(simp)
  4037 apply(simp add: abs_fresh)
  4045     apply(simp add: abs_fresh)
  4038 apply(auto)[1]
  4046    apply(auto)[1]
  4039 apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
  4047     apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
  4040 apply(simp add: calc_atm fresh_atm fresh_prod)
  4048     apply(simp add: calc_atm fresh_atm fresh_prod)
  4041 apply(subgoal_tac"b\<sharp>([(c,a)]\<bullet>M)")
  4049    apply(subgoal_tac"b\<sharp>([(c,a)]\<bullet>M)")
  4042 apply(simp add: forget)
  4050     apply(simp add: forget)
  4043 apply(simp add: trm.inject)
  4051     apply(simp add: trm.inject)
  4044 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4052     apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4045 apply(perm_simp)
  4053     apply(perm_simp)
  4046 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
  4054    apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
  4047 apply(simp add: trm.inject)
  4055   apply(simp add: trm.inject)
  4048 apply(rule conjI)
  4056   apply(rule conjI)
  4049 apply(rule sym)
  4057    apply(rule sym)
  4050 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4058    apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4051 apply(rule sym)
  4059   apply(rule sym)
  4052 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4060   apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4053 done
  4061   done
  4054 
  4062 
  4055 lemma better_Cut_substn1:
  4063 lemma better_Cut_substn1:
  4056   assumes a: "y\<sharp>(x,N)" "a\<sharp>(b,P)" "M\<noteq>Ax y a"
  4064   assumes a: "y\<sharp>(x,N)" "a\<sharp>(b,P)" "M\<noteq>Ax y a"
  4057   shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.(M{y:=<b>.P}) (x).N"
  4065   shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.(M{y:=<b>.P}) (x).N"
  4058 using a
  4066   using a
  4059 apply -
  4067   apply -
  4060 apply(generate_fresh "coname")
  4068   apply(generate_fresh "coname")
  4061 apply(generate_fresh "name")
  4069   apply(generate_fresh "name")
  4062 apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4070   apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4063 apply(simp)
  4071    apply(simp)
  4064 apply(rule trans)
  4072    apply(rule trans)
  4065 apply(rule better_Cut_substn)
  4073     apply(rule better_Cut_substn)
  4066 apply(simp add: abs_fresh)
  4074      apply(simp add: abs_fresh)
  4067 apply(simp add: abs_fresh)
  4075     apply(simp add: abs_fresh)
  4068 apply(auto)[1]
  4076    apply(auto)[1]
  4069 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
  4077     apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
  4070 apply(simp add: calc_atm fresh_atm fresh_prod)
  4078     apply(simp add: calc_atm fresh_atm fresh_prod)
  4071 apply(subgoal_tac"y\<sharp>([(ca,x)]\<bullet>N)")
  4079    apply(subgoal_tac"y\<sharp>([(ca,x)]\<bullet>N)")
  4072 apply(simp add: forget)
  4080     apply(simp add: forget)
  4073 apply(simp add: trm.inject)
  4081     apply(simp add: trm.inject)
  4074 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4082     apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4075 apply(perm_simp)
  4083     apply(perm_simp)
  4076 apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
  4084    apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
  4077 apply(simp add: trm.inject)
  4085   apply(simp add: trm.inject)
  4078 apply(rule conjI)
  4086   apply(rule conjI)
  4079 apply(rule sym)
  4087    apply(rule sym)
  4080 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4088    apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4081 apply(rule sym)
  4089   apply(rule sym)
  4082 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4090   apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4083 done
  4091   done
  4084 
  4092 
  4085 lemma better_Cut_substn2:
  4093 lemma better_Cut_substn2:
  4086   assumes a: "x\<sharp>(P,y)" "y\<sharp>M" 
  4094   assumes a: "x\<sharp>(P,y)" "y\<sharp>M" 
  4087   shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.M (x).(N{y:=<b>.P})"
  4095   shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.M (x).(N{y:=<b>.P})"
  4088 using a
  4096   using a
  4089 apply -
  4097   apply -
  4090 apply(generate_fresh "coname")
  4098   apply(generate_fresh "coname")
  4091 apply(generate_fresh "name")
  4099   apply(generate_fresh "name")
  4092 apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4100   apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
  4093 apply(simp)
  4101    apply(simp)
  4094 apply(rule trans)
  4102    apply(rule trans)
  4095 apply(rule better_Cut_substn)
  4103     apply(rule better_Cut_substn)
  4096 apply(simp add: abs_fresh)
  4104      apply(simp add: abs_fresh)
  4097 apply(simp add: abs_fresh)
  4105     apply(simp add: abs_fresh)
  4098 apply(auto)[1]
  4106    apply(auto)[1]
  4099 apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
  4107     apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
  4100 apply(simp add: calc_atm fresh_atm)
  4108     apply(simp add: calc_atm fresh_atm)
  4101 apply(subgoal_tac"y\<sharp>([(c,a)]\<bullet>M)")
  4109    apply(subgoal_tac"y\<sharp>([(c,a)]\<bullet>M)")
  4102 apply(simp add: forget)
  4110     apply(simp add: forget)
  4103 apply(simp add: trm.inject)
  4111     apply(simp add: trm.inject)
  4104 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4112     apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4105 apply(perm_simp)
  4113     apply(perm_simp)
  4106 apply(simp add: fresh_left calc_atm)
  4114    apply(simp add: fresh_left calc_atm)
  4107 apply(simp add: trm.inject)
  4115   apply(simp add: trm.inject)
  4108 apply(rule conjI)
  4116   apply(rule conjI)
  4109 apply(rule sym)
  4117    apply(rule sym)
  4110 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4118    apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4111 apply(rule sym)
  4119   apply(rule sym)
  4112 apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4120   apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
  4113 done
  4121   done
  4114 
  4122 
  4115 lemma psubst_fresh_name:
  4123 lemma psubst_fresh_name:
  4116   fixes x::"name"
  4124   fixes x::"name"
  4117   assumes a: "x\<sharp>\<theta>_n" "x\<sharp>\<theta>_c" "x\<sharp>M"
  4125   assumes a: "x\<sharp>\<theta>_n" "x\<sharp>\<theta>_c" "x\<sharp>M"
  4118   shows "x\<sharp>\<theta>_n,\<theta>_c<M>"
  4126   shows "x\<sharp>\<theta>_n,\<theta>_c<M>"
  4119 using a
  4127   using a
  4120 apply(nominal_induct M avoiding: x \<theta>_n \<theta>_c rule: trm.strong_induct)
  4128   apply(nominal_induct M avoiding: x \<theta>_n \<theta>_c rule: trm.strong_induct)
  4121 apply(simp add: lookup_freshness)
  4129              apply(simp add: lookup_freshness)
  4122 apply(auto simp add: abs_fresh)[1]
  4130             apply(auto simp add: abs_fresh)[1]
  4123 apply(simp add: lookupc_freshness)
  4131                  apply(simp add: lookupc_freshness)
  4124 apply(simp add: lookupc_freshness)
  4132                 apply(simp add: lookupc_freshness)
  4125 apply(simp add: lookupc_freshness)
  4133                apply(simp add: lookupc_freshness)
  4126 apply(simp add: lookupd_freshness)
  4134               apply(simp add: lookupd_freshness)
  4127 apply(simp add: lookupd_freshness)
  4135              apply(simp add: lookupd_freshness)
  4128 apply(simp add: lookupc_freshness)
  4136             apply(simp add: lookupc_freshness)
  4129 apply(simp)
  4137            apply(simp)
  4130 apply(case_tac "findc \<theta>_c coname")
  4138            apply(case_tac "findc \<theta>_c coname")
  4131 apply(auto simp add: abs_fresh)[1]
  4139             apply(auto simp add: abs_fresh)[1]
  4132 apply(auto)[1]
  4140            apply(auto)[1]
  4133 apply(generate_fresh "coname")
  4141            apply(generate_fresh "coname")
  4134 apply(fresh_fun_simp)
  4142            apply(fresh_fun_simp)
  4135 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4143            apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4136 apply(simp)
  4144           apply(simp)
  4137 apply(case_tac "findn \<theta>_n name")
  4145           apply(case_tac "findn \<theta>_n name")
  4138 apply(auto simp add: abs_fresh)[1]
  4146            apply(auto simp add: abs_fresh)[1]
  4139 apply(auto)[1]
  4147           apply(auto)[1]
  4140 apply(generate_fresh "name")
  4148           apply(generate_fresh "name")
  4141 apply(fresh_fun_simp)
  4149           apply(fresh_fun_simp)
  4142 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4150           apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4143 apply(simp)
  4151          apply(simp)
  4144 apply(case_tac "findc \<theta>_c coname3")
  4152          apply(case_tac "findc \<theta>_c coname3")
  4145 apply(auto simp add: abs_fresh)[1]
  4153           apply(auto simp add: abs_fresh)[1]
  4146 apply(auto)[1]
  4154          apply(auto)[1]
  4147 apply(generate_fresh "coname")
  4155          apply(generate_fresh "coname")
  4148 apply(fresh_fun_simp)
  4156          apply(fresh_fun_simp)
  4149 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4157          apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4150 apply(simp)
  4158         apply(simp)
  4151 apply(case_tac "findn \<theta>_n name2")
  4159         apply(case_tac "findn \<theta>_n name2")
  4152 apply(auto simp add: abs_fresh)[1]
  4160          apply(auto simp add: abs_fresh)[1]
  4153 apply(auto)[1]
  4161         apply(auto)[1]
  4154 apply(generate_fresh "name")
  4162         apply(generate_fresh "name")
  4155 apply(fresh_fun_simp)
  4163         apply(fresh_fun_simp)
  4156 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4164         apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4157 apply(simp)
  4165        apply(simp)
  4158 apply(case_tac "findn \<theta>_n name2")
  4166        apply(case_tac "findn \<theta>_n name2")
  4159 apply(auto simp add: abs_fresh)[1]
  4167         apply(auto simp add: abs_fresh)[1]
  4160 apply(auto)[1]
  4168        apply(auto)[1]
  4161 apply(generate_fresh "name")
  4169        apply(generate_fresh "name")
  4162 apply(fresh_fun_simp)
  4170        apply(fresh_fun_simp)
  4163 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4171        apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4164 apply(simp)
  4172       apply(simp)
  4165 apply(case_tac "findc \<theta>_c coname2")
  4173       apply(case_tac "findc \<theta>_c coname2")
  4166 apply(auto simp add: abs_fresh)[1]
  4174        apply(auto simp add: abs_fresh)[1]
  4167 apply(auto)[1]
  4175       apply(auto)[1]
  4168 apply(generate_fresh "coname")
  4176       apply(generate_fresh "coname")
  4169 apply(fresh_fun_simp)
  4177       apply(fresh_fun_simp)
  4170 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4178       apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4171 apply(simp)
  4179      apply(simp)
  4172 apply(case_tac "findc \<theta>_c coname2")
  4180      apply(case_tac "findc \<theta>_c coname2")
  4173 apply(auto simp add: abs_fresh)[1]
  4181       apply(auto simp add: abs_fresh)[1]
  4174 apply(auto)[1]
  4182      apply(auto)[1]
  4175 apply(generate_fresh "coname")
  4183      apply(generate_fresh "coname")
  4176 apply(fresh_fun_simp)
  4184      apply(fresh_fun_simp)
  4177 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4185      apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4178 apply(simp)
  4186     apply(simp)
  4179 apply(case_tac "findn \<theta>_n name3")
  4187     apply(case_tac "findn \<theta>_n name3")
  4180 apply(auto simp add: abs_fresh)[1]
  4188      apply(auto simp add: abs_fresh)[1]
  4181 apply(auto)[1]
  4189     apply(auto)[1]
  4182 apply(generate_fresh "name")
  4190     apply(generate_fresh "name")
  4183 apply(fresh_fun_simp)
  4191     apply(fresh_fun_simp)
  4184 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4192     apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4185 apply(simp)
  4193    apply(simp)
  4186 apply(case_tac "findc \<theta>_c coname2")
  4194    apply(case_tac "findc \<theta>_c coname2")
  4187 apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
  4195     apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
  4188 apply(auto)[1]
  4196    apply(auto)[1]
  4189 apply(generate_fresh "coname")
  4197    apply(generate_fresh "coname")
  4190 apply(fresh_fun_simp)
  4198    apply(fresh_fun_simp)
  4191 apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
  4199    apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
  4192 apply(simp)
  4200   apply(simp)
  4193 apply(case_tac "findn \<theta>_n name2")
  4201   apply(case_tac "findn \<theta>_n name2")
  4194 apply(auto simp add: abs_fresh)[1]
  4202    apply(auto simp add: abs_fresh)[1]
  4195 apply(auto)[1]
  4203   apply(auto)[1]
  4196 apply(generate_fresh "name")
  4204   apply(generate_fresh "name")
  4197 apply(fresh_fun_simp)
  4205   apply(fresh_fun_simp)
  4198 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4206   apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4199 done
  4207   done
  4200 
  4208 
  4201 lemma psubst_fresh_coname:
  4209 lemma psubst_fresh_coname:
  4202   fixes a::"coname"
  4210   fixes a::"coname"
  4203   assumes a: "a\<sharp>\<theta>_n" "a\<sharp>\<theta>_c" "a\<sharp>M"
  4211   assumes a: "a\<sharp>\<theta>_n" "a\<sharp>\<theta>_c" "a\<sharp>M"
  4204   shows "a\<sharp>\<theta>_n,\<theta>_c<M>"
  4212   shows "a\<sharp>\<theta>_n,\<theta>_c<M>"
  4205 using a
  4213   using a
  4206 apply(nominal_induct M avoiding: a \<theta>_n \<theta>_c rule: trm.strong_induct)
  4214   apply(nominal_induct M avoiding: a \<theta>_n \<theta>_c rule: trm.strong_induct)
  4207 apply(simp add: lookup_freshness)
  4215              apply(simp add: lookup_freshness)
  4208 apply(auto simp add: abs_fresh)[1]
  4216             apply(auto simp add: abs_fresh)[1]
  4209 apply(simp add: lookupd_freshness)
  4217                  apply(simp add: lookupd_freshness)
  4210 apply(simp add: lookupd_freshness)
  4218                 apply(simp add: lookupd_freshness)
  4211 apply(simp add: lookupc_freshness)
  4219                apply(simp add: lookupc_freshness)
  4212 apply(simp add: lookupd_freshness)
  4220               apply(simp add: lookupd_freshness)
  4213 apply(simp add: lookupc_freshness)
  4221              apply(simp add: lookupc_freshness)
  4214 apply(simp add: lookupd_freshness)
  4222             apply(simp add: lookupd_freshness)
  4215 apply(simp)
  4223            apply(simp)
  4216 apply(case_tac "findc \<theta>_c coname")
  4224            apply(case_tac "findc \<theta>_c coname")
  4217 apply(auto simp add: abs_fresh)[1]
  4225             apply(auto simp add: abs_fresh)[1]
  4218 apply(auto)[1]
  4226            apply(auto)[1]
  4219 apply(generate_fresh "coname")
  4227            apply(generate_fresh "coname")
  4220 apply(fresh_fun_simp)
  4228            apply(fresh_fun_simp)
  4221 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4229            apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4222 apply(simp)
  4230           apply(simp)
  4223 apply(case_tac "findn \<theta>_n name")
  4231           apply(case_tac "findn \<theta>_n name")
  4224 apply(auto simp add: abs_fresh)[1]
  4232            apply(auto simp add: abs_fresh)[1]
  4225 apply(auto)[1]
  4233           apply(auto)[1]
  4226 apply(generate_fresh "name")
  4234           apply(generate_fresh "name")
  4227 apply(fresh_fun_simp)
  4235           apply(fresh_fun_simp)
  4228 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4236           apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4229 apply(simp)
  4237          apply(simp)
  4230 apply(case_tac "findc \<theta>_c coname3")
  4238          apply(case_tac "findc \<theta>_c coname3")
  4231 apply(auto simp add: abs_fresh)[1]
  4239           apply(auto simp add: abs_fresh)[1]
  4232 apply(auto)[1]
  4240          apply(auto)[1]
  4233 apply(generate_fresh "coname")
  4241          apply(generate_fresh "coname")
  4234 apply(fresh_fun_simp)
  4242          apply(fresh_fun_simp)
  4235 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4243          apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4236 apply(simp)
  4244         apply(simp)
  4237 apply(case_tac "findn \<theta>_n name2")
  4245         apply(case_tac "findn \<theta>_n name2")
  4238 apply(auto simp add: abs_fresh)[1]
  4246          apply(auto simp add: abs_fresh)[1]
  4239 apply(auto)[1]
  4247         apply(auto)[1]
  4240 apply(generate_fresh "name")
  4248         apply(generate_fresh "name")
  4241 apply(fresh_fun_simp)
  4249         apply(fresh_fun_simp)
  4242 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4250         apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4243 apply(simp)
  4251        apply(simp)
  4244 apply(case_tac "findn \<theta>_n name2")
  4252        apply(case_tac "findn \<theta>_n name2")
  4245 apply(auto simp add: abs_fresh)[1]
  4253         apply(auto simp add: abs_fresh)[1]
  4246 apply(auto)[1]
  4254        apply(auto)[1]
  4247 apply(generate_fresh "name")
  4255        apply(generate_fresh "name")
  4248 apply(fresh_fun_simp)
  4256        apply(fresh_fun_simp)
  4249 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4257        apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4250 apply(simp)
  4258       apply(simp)
  4251 apply(case_tac "findc \<theta>_c coname2")
  4259       apply(case_tac "findc \<theta>_c coname2")
  4252 apply(auto simp add: abs_fresh)[1]
  4260        apply(auto simp add: abs_fresh)[1]
  4253 apply(auto)[1]
  4261       apply(auto)[1]
  4254 apply(generate_fresh "coname")
  4262       apply(generate_fresh "coname")
  4255 apply(fresh_fun_simp)
  4263       apply(fresh_fun_simp)
  4256 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4264       apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4257 apply(simp)
  4265      apply(simp)
  4258 apply(case_tac "findc \<theta>_c coname2")
  4266      apply(case_tac "findc \<theta>_c coname2")
  4259 apply(auto simp add: abs_fresh)[1]
  4267       apply(auto simp add: abs_fresh)[1]
  4260 apply(auto)[1]
  4268      apply(auto)[1]
  4261 apply(generate_fresh "coname")
  4269      apply(generate_fresh "coname")
  4262 apply(fresh_fun_simp)
  4270      apply(fresh_fun_simp)
  4263 apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4271      apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
  4264 apply(simp)
  4272     apply(simp)
  4265 apply(case_tac "findn \<theta>_n name3")
  4273     apply(case_tac "findn \<theta>_n name3")
  4266 apply(auto simp add: abs_fresh)[1]
  4274      apply(auto simp add: abs_fresh)[1]
  4267 apply(auto)[1]
  4275     apply(auto)[1]
  4268 apply(generate_fresh "name")
  4276     apply(generate_fresh "name")
  4269 apply(fresh_fun_simp)
  4277     apply(fresh_fun_simp)
  4270 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4278     apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4271 apply(simp)
  4279    apply(simp)
  4272 apply(case_tac "findc \<theta>_c coname2")
  4280    apply(case_tac "findc \<theta>_c coname2")
  4273 apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
  4281     apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
  4274 apply(auto)[1]
  4282    apply(auto)[1]
  4275 apply(generate_fresh "coname")
  4283    apply(generate_fresh "coname")
  4276 apply(fresh_fun_simp)
  4284    apply(fresh_fun_simp)
  4277 apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
  4285    apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
  4278 apply(simp)
  4286   apply(simp)
  4279 apply(case_tac "findn \<theta>_n name2")
  4287   apply(case_tac "findn \<theta>_n name2")
  4280 apply(auto simp add: abs_fresh)[1]
  4288    apply(auto simp add: abs_fresh)[1]
  4281 apply(auto)[1]
  4289   apply(auto)[1]
  4282 apply(generate_fresh "name")
  4290   apply(generate_fresh "name")
  4283 apply(fresh_fun_simp)
  4291   apply(fresh_fun_simp)
  4284 apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4292   apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
  4285 done
  4293   done
  4286 
  4294 
  4287 lemma psubst_csubst:
  4295 lemma psubst_csubst:
  4288   assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
  4296   assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
  4289   shows "\<theta>_n,((a,x,P)#\<theta>_c)<M> = ((\<theta>_n,\<theta>_c<M>){a:=(x).P})"
  4297   shows "\<theta>_n,((a,x,P)#\<theta>_c)<M> = ((\<theta>_n,\<theta>_c<M>){a:=(x).P})"
  4290 using a
  4298   using a
  4291 apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
  4299   apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
  4292 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4300              apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4293 apply(simp add: lookup_csubst)
  4301              apply(simp add: lookup_csubst)
  4294 apply(simp add: fresh_list_cons fresh_prod)
  4302             apply(simp add: fresh_list_cons fresh_prod)
  4295 apply(auto)[1]
  4303             apply(auto)[1]
  4296 apply(rule sym)
  4304                  apply(rule sym)
  4297 apply(rule trans)
  4305                  apply(rule trans)
  4298 apply(rule better_Cut_substc)
  4306                   apply(rule better_Cut_substc)
  4299 apply(simp)
  4307                    apply(simp)
  4300 apply(simp add: abs_fresh fresh_atm)
  4308                   apply(simp add: abs_fresh fresh_atm)
  4301 apply(simp add: lookupd_fresh)
  4309                  apply(simp add: lookupd_fresh)
  4302 apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
  4310                  apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
  4303 apply(simp add: forget)
  4311                   apply(simp add: forget)
  4304 apply(simp add: trm.inject)
  4312                   apply(simp add: trm.inject)
  4305 apply(rule sym)
  4313                   apply(rule sym)
  4306 apply(simp add: alpha nrename_swap fresh_atm)
  4314                   apply(simp add: alpha nrename_swap fresh_atm)
  4307 apply(rule lookupc_freshness)
  4315                  apply(rule lookupc_freshness)
  4308 apply(simp add: fresh_atm)
  4316                  apply(simp add: fresh_atm)
  4309 apply(rule sym)
  4317                 apply(rule sym)
  4310 apply(rule trans)
  4318                 apply(rule trans)
  4311 apply(rule better_Cut_substc)
  4319                  apply(rule better_Cut_substc)
  4312 apply(simp)
  4320                   apply(simp)
  4313 apply(simp add: abs_fresh fresh_atm)
  4321                  apply(simp add: abs_fresh fresh_atm)
  4314 apply(simp)
  4322                 apply(simp)
  4315 apply(rule conjI)
  4323                 apply(rule conjI)
  4316 apply(rule impI)
  4324                  apply(rule impI)
  4317 apply(simp add: lookupd_unicity)
  4325                  apply(simp add: lookupd_unicity)
  4318 apply(rule impI)
  4326                 apply(rule impI)
  4319 apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
  4327                 apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
  4320 apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
  4328                  apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
  4321 apply(simp add: forget)
  4329                   apply(simp add: forget)
  4322 apply(rule lookupd_freshness)
  4330                  apply(rule lookupd_freshness)
  4323 apply(simp add: fresh_atm)
  4331                  apply(simp add: fresh_atm)
  4324 apply(rule lookupc_freshness)
  4332                 apply(rule lookupc_freshness)
  4325 apply(simp add: fresh_atm)
  4333                 apply(simp add: fresh_atm)
  4326 apply(rule sym)
  4334                apply(rule sym)
  4327 apply(rule trans)
  4335                apply(rule trans)
  4328 apply(rule better_Cut_substc)
  4336                 apply(rule better_Cut_substc)
  4329 apply(simp)
  4337                  apply(simp)
  4330 apply(simp add: abs_fresh fresh_atm)
  4338                 apply(simp add: abs_fresh fresh_atm)
  4331 apply(simp)
  4339                apply(simp)
  4332 apply(rule conjI)
  4340                apply(rule conjI)
  4333 apply(rule impI)
  4341                 apply(rule impI)
  4334 apply(drule ax_psubst)
  4342                 apply(drule ax_psubst)
  4335 apply(simp)
  4343                   apply(simp)
  4336 apply(simp)
  4344                  apply(simp)
  4337 apply(blast)
  4345                 apply(blast)
  4338 apply(rule impI)
  4346                apply(rule impI)
  4339 apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
  4347                apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
  4340 apply(simp add: forget)
  4348                 apply(simp add: forget)
  4341 apply(rule lookupc_freshness)
  4349                apply(rule lookupc_freshness)
  4342 apply(simp add: fresh_atm)
  4350                apply(simp add: fresh_atm)
  4343 apply(rule sym)
  4351               apply(rule sym)
  4344 apply(rule trans)
  4352               apply(rule trans)
  4345 apply(rule better_Cut_substc)
  4353                apply(rule better_Cut_substc)
  4346 apply(simp)
  4354                 apply(simp)
  4347 apply(simp add: abs_fresh fresh_atm)
  4355                apply(simp add: abs_fresh fresh_atm)
  4348 apply(simp)
  4356               apply(simp)
  4349 apply(rule conjI)
  4357               apply(rule conjI)
  4350 apply(rule impI)
  4358                apply(rule impI)
  4351 apply(simp add: trm.inject)
  4359                apply(simp add: trm.inject)
  4352 apply(rule sym)
  4360                apply(rule sym)
  4353 apply(simp add: alpha)
  4361                apply(simp add: alpha)
  4354 apply(simp add: alpha nrename_swap fresh_atm)
  4362                apply(simp add: alpha nrename_swap fresh_atm)
  4355 apply(simp add: lookupd_fresh)
  4363               apply(simp add: lookupd_fresh)
  4356 apply(rule sym)
  4364              apply(rule sym)
  4357 apply(rule trans)
  4365              apply(rule trans)
  4358 apply(rule better_Cut_substc)
  4366               apply(rule better_Cut_substc)
  4359 apply(simp)
  4367                apply(simp)
  4360 apply(simp add: abs_fresh fresh_atm)
  4368               apply(simp add: abs_fresh fresh_atm)
  4361 apply(simp)
  4369              apply(simp)
  4362 apply(rule conjI)
  4370              apply(rule conjI)
  4363 apply(rule impI)
  4371               apply(rule impI)
  4364 apply(simp add: lookupd_unicity)
  4372               apply(simp add: lookupd_unicity)
  4365 apply(rule impI)
  4373              apply(rule impI)
  4366 apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
  4374              apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
  4367 apply(simp add: forget)
  4375               apply(simp add: forget)
  4368 apply(rule lookupd_freshness)
  4376              apply(rule lookupd_freshness)
  4369 apply(simp add: fresh_atm)
  4377              apply(simp add: fresh_atm)
  4370 apply(rule sym)
  4378             apply(rule sym)
  4371 apply(rule trans)
  4379             apply(rule trans)
  4372 apply(rule better_Cut_substc)
  4380              apply(rule better_Cut_substc)
  4373 apply(simp)
  4381               apply(simp)
  4374 apply(simp add: abs_fresh fresh_atm)
  4382              apply(simp add: abs_fresh fresh_atm)
  4375 apply(simp)
  4383             apply(simp)
  4376 apply(rule impI)
  4384             apply(rule impI)
  4377 apply(drule ax_psubst)
  4385             apply(drule ax_psubst)
  4378 apply(simp)
  4386               apply(simp)
  4379 apply(simp)
  4387              apply(simp)
  4380 apply(blast)
  4388             apply(blast)
  4381 (* NotR *)
  4389     (* NotR *)
  4382 apply(simp)
  4390            apply(simp)
  4383 apply(case_tac "findc \<theta>_c coname")
  4391            apply(case_tac "findc \<theta>_c coname")
  4384 apply(simp)
  4392             apply(simp)
  4385 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4393             apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4386 apply(simp)
  4394            apply(simp)
  4387 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4395            apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4388 apply(drule cmaps_false)
  4396             apply(drule cmaps_false)
  4389 apply(assumption)
  4397              apply(assumption)
  4390 apply(simp)
  4398             apply(simp)
  4391 apply(generate_fresh "coname")
  4399            apply(generate_fresh "coname")
  4392 apply(fresh_fun_simp)
  4400            apply(fresh_fun_simp)
  4393 apply(fresh_fun_simp)
  4401            apply(fresh_fun_simp)
  4394 apply(rule sym)
  4402            apply(rule sym)
  4395 apply(rule trans)
  4403            apply(rule trans)
  4396 apply(rule better_Cut_substc1)
  4404             apply(rule better_Cut_substc1)
  4397 apply(simp)
  4405              apply(simp)
  4398 apply(simp add: cmaps_fresh)
  4406             apply(simp add: cmaps_fresh)
  4399 apply(auto simp add: fresh_prod fresh_atm)[1]
  4407            apply(auto simp add: fresh_prod fresh_atm)[1]
  4400 (* NotL *)
  4408     (* NotL *)
  4401 apply(simp)
  4409           apply(simp)
  4402 apply(case_tac "findn \<theta>_n name")
  4410           apply(case_tac "findn \<theta>_n name")
  4403 apply(simp)
  4411            apply(simp)
  4404 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4412            apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4405 apply(simp)
  4413           apply(simp)
  4406 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4414           apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4407 apply(generate_fresh "name")
  4415           apply(generate_fresh "name")
  4408 apply(fresh_fun_simp)
  4416           apply(fresh_fun_simp)
  4409 apply(fresh_fun_simp)
  4417           apply(fresh_fun_simp)
  4410 apply(drule_tac a="a" in nmaps_fresh)
  4418           apply(drule_tac a="a" in nmaps_fresh)
  4411 apply(assumption)
  4419            apply(assumption)
  4412 apply(rule sym)
  4420           apply(rule sym)
  4413 apply(rule trans)
  4421           apply(rule trans)
  4414 apply(rule better_Cut_substc2)
  4422            apply(rule better_Cut_substc2)
  4415 apply(simp)
  4423              apply(simp)
  4416 apply(simp)
  4424             apply(simp)
  4417 apply(simp)
  4425            apply(simp)
  4418 apply(simp)
  4426           apply(simp)
  4419 (* AndR *)
  4427     (* AndR *)
  4420 apply(simp)
  4428          apply(simp)
  4421 apply(case_tac "findc \<theta>_c coname3")
  4429          apply(case_tac "findc \<theta>_c coname3")
  4422 apply(simp)
  4430           apply(simp)
  4423 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4431           apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4424 apply(simp)
  4432          apply(simp)
  4425 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4433          apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4426 apply(drule cmaps_false)
  4434           apply(drule cmaps_false)
  4427 apply(assumption)
  4435            apply(assumption)
  4428 apply(simp)
  4436           apply(simp)
  4429 apply(generate_fresh "coname")
  4437          apply(generate_fresh "coname")
  4430 apply(fresh_fun_simp)
  4438          apply(fresh_fun_simp)
  4431 apply(fresh_fun_simp)
  4439          apply(fresh_fun_simp)
  4432 apply(rule sym)
  4440          apply(rule sym)
  4433 apply(rule trans)
  4441          apply(rule trans)
  4434 apply(rule better_Cut_substc1)
  4442           apply(rule better_Cut_substc1)
  4435 apply(simp)
  4443            apply(simp)
  4436 apply(simp add: cmaps_fresh)
  4444           apply(simp add: cmaps_fresh)
  4437 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4445          apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4438 (* AndL1 *)
  4446     (* AndL1 *)
  4439 apply(simp)
  4447         apply(simp)
  4440 apply(case_tac "findn \<theta>_n name2")
  4448         apply(case_tac "findn \<theta>_n name2")
  4441 apply(simp)
  4449          apply(simp)
  4442 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4450          apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4443 apply(simp)
  4451         apply(simp)
  4444 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4452         apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4445 apply(generate_fresh "name")
  4453         apply(generate_fresh "name")
  4446 apply(fresh_fun_simp)
  4454         apply(fresh_fun_simp)
  4447 apply(fresh_fun_simp)
  4455         apply(fresh_fun_simp)
  4448 apply(drule_tac a="a" in nmaps_fresh)
  4456         apply(drule_tac a="a" in nmaps_fresh)
  4449 apply(assumption)
  4457          apply(assumption)
  4450 apply(rule sym)
  4458         apply(rule sym)
  4451 apply(rule trans)
  4459         apply(rule trans)
  4452 apply(rule better_Cut_substc2)
  4460          apply(rule better_Cut_substc2)
  4453 apply(simp)
  4461            apply(simp)
  4454 apply(simp)
  4462           apply(simp)
  4455 apply(simp)
  4463          apply(simp)
  4456 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4464         apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4457 (* AndL2 *)
  4465     (* AndL2 *)
  4458 apply(simp)
  4466        apply(simp)
  4459 apply(case_tac "findn \<theta>_n name2")
  4467        apply(case_tac "findn \<theta>_n name2")
  4460 apply(simp)
  4468         apply(simp)
  4461 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4469         apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4462 apply(simp)
  4470        apply(simp)
  4463 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4471        apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4464 apply(generate_fresh "name")
  4472        apply(generate_fresh "name")
  4465 apply(fresh_fun_simp)
  4473        apply(fresh_fun_simp)
  4466 apply(fresh_fun_simp)
  4474        apply(fresh_fun_simp)
  4467 apply(drule_tac a="a" in nmaps_fresh)
  4475        apply(drule_tac a="a" in nmaps_fresh)
  4468 apply(assumption)
  4476         apply(assumption)
  4469 apply(rule sym)
  4477        apply(rule sym)
  4470 apply(rule trans)
  4478        apply(rule trans)
  4471 apply(rule better_Cut_substc2)
  4479         apply(rule better_Cut_substc2)
  4472 apply(simp)
  4480           apply(simp)
  4473 apply(simp)
  4481          apply(simp)
  4474 apply(simp)
  4482         apply(simp)
  4475 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4483        apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4476 (* OrR1 *)
  4484     (* OrR1 *)
  4477 apply(simp)
  4485       apply(simp)
  4478 apply(case_tac "findc \<theta>_c coname2")
  4486       apply(case_tac "findc \<theta>_c coname2")
  4479 apply(simp)
  4487        apply(simp)
  4480 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4488        apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4481 apply(simp)
  4489       apply(simp)
  4482 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4490       apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4483 apply(drule cmaps_false)
  4491        apply(drule cmaps_false)
  4484 apply(assumption)
  4492         apply(assumption)
  4485 apply(simp)
  4493        apply(simp)
  4486 apply(generate_fresh "coname")
  4494       apply(generate_fresh "coname")
  4487 apply(fresh_fun_simp)
  4495       apply(fresh_fun_simp)
  4488 apply(fresh_fun_simp)
  4496       apply(fresh_fun_simp)
  4489 apply(rule sym)
  4497       apply(rule sym)
  4490 apply(rule trans)
  4498       apply(rule trans)
  4491 apply(rule better_Cut_substc1)
  4499        apply(rule better_Cut_substc1)
  4492 apply(simp)
  4500         apply(simp)
  4493 apply(simp add: cmaps_fresh)
  4501        apply(simp add: cmaps_fresh)
  4494 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4502       apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4495 (* OrR2 *)
  4503     (* OrR2 *)
  4496 apply(simp)
  4504      apply(simp)
  4497 apply(case_tac "findc \<theta>_c coname2")
  4505      apply(case_tac "findc \<theta>_c coname2")
  4498 apply(simp)
  4506       apply(simp)
  4499 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4507       apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4500 apply(simp)
  4508      apply(simp)
  4501 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4509      apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4502 apply(drule cmaps_false)
  4510       apply(drule cmaps_false)
  4503 apply(assumption)
  4511        apply(assumption)
  4504 apply(simp)
  4512       apply(simp)
  4505 apply(generate_fresh "coname")
  4513      apply(generate_fresh "coname")
  4506 apply(fresh_fun_simp)
  4514      apply(fresh_fun_simp)
  4507 apply(fresh_fun_simp)
  4515      apply(fresh_fun_simp)
  4508 apply(rule sym)
  4516      apply(rule sym)
  4509 apply(rule trans)
  4517      apply(rule trans)
  4510 apply(rule better_Cut_substc1)
  4518       apply(rule better_Cut_substc1)
  4511 apply(simp)
  4519        apply(simp)
  4512 apply(simp add: cmaps_fresh)
  4520       apply(simp add: cmaps_fresh)
  4513 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4521      apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4514 (* OrL *)
  4522     (* OrL *)
  4515 apply(simp)
  4523     apply(simp)
  4516 apply(case_tac "findn \<theta>_n name3")
  4524     apply(case_tac "findn \<theta>_n name3")
  4517 apply(simp)
  4525      apply(simp)
  4518 apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
  4526      apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
  4519 apply(simp)
  4527     apply(simp)
  4520 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4528     apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4521 apply(generate_fresh "name")
  4529     apply(generate_fresh "name")
  4522 apply(fresh_fun_simp)
  4530     apply(fresh_fun_simp)
  4523 apply(fresh_fun_simp)
  4531     apply(fresh_fun_simp)
  4524 apply(drule_tac a="a" in nmaps_fresh)
  4532     apply(drule_tac a="a" in nmaps_fresh)
  4525 apply(assumption)
  4533      apply(assumption)
  4526 apply(rule sym)
  4534     apply(rule sym)
  4527 apply(rule trans)
  4535     apply(rule trans)
  4528 apply(rule better_Cut_substc2)
  4536      apply(rule better_Cut_substc2)
  4529 apply(simp)
  4537        apply(simp)
  4530 apply(simp)
  4538       apply(simp)
  4531 apply(simp)
  4539      apply(simp)
  4532 apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
  4540     apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
  4533 (* ImpR *)
  4541     (* ImpR *)
  4534 apply(simp)
  4542    apply(simp)
  4535 apply(case_tac "findc \<theta>_c coname2")
  4543    apply(case_tac "findc \<theta>_c coname2")
  4536 apply(simp)
  4544     apply(simp)
  4537 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4545     apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4538 apply(simp)
  4546    apply(simp)
  4539 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4547    apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4540 apply(drule cmaps_false)
  4548     apply(drule cmaps_false)
  4541 apply(assumption)
  4549      apply(assumption)
  4542 apply(simp)
  4550     apply(simp)
  4543 apply(generate_fresh "coname")
  4551    apply(generate_fresh "coname")
  4544 apply(fresh_fun_simp)
  4552    apply(fresh_fun_simp)
  4545 apply(fresh_fun_simp)
  4553    apply(fresh_fun_simp)
  4546 apply(rule sym)
  4554    apply(rule sym)
  4547 apply(rule trans)
  4555    apply(rule trans)
  4548 apply(rule better_Cut_substc1)
  4556     apply(rule better_Cut_substc1)
  4549 apply(simp)
  4557      apply(simp)
  4550 apply(simp add: cmaps_fresh)
  4558     apply(simp add: cmaps_fresh)
  4551 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4559    apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4552 (* ImpL *)
  4560     (* ImpL *)
  4553 apply(simp)
  4561   apply(simp)
  4554 apply(case_tac "findn \<theta>_n name2")
  4562   apply(case_tac "findn \<theta>_n name2")
  4555 apply(simp)
  4563    apply(simp)
  4556 apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
  4564    apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
  4557 apply(simp)
  4565   apply(simp)
  4558 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4566   apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4559 apply(generate_fresh "name")
  4567   apply(generate_fresh "name")
  4560 apply(fresh_fun_simp)
  4568   apply(fresh_fun_simp)
  4561 apply(fresh_fun_simp)
  4569   apply(fresh_fun_simp)
  4562 apply(simp add: abs_fresh subst_fresh)
  4570   apply(simp add: abs_fresh subst_fresh)
  4563 apply(drule_tac a="a" in nmaps_fresh)
  4571   apply(drule_tac a="a" in nmaps_fresh)
  4564 apply(assumption)
  4572    apply(assumption)
  4565 apply(rule sym)
  4573   apply(rule sym)
  4566 apply(rule trans)
  4574   apply(rule trans)
  4567 apply(rule better_Cut_substc2)
  4575    apply(rule better_Cut_substc2)
  4568 apply(simp)
  4576      apply(simp)
  4569 apply(simp)
  4577     apply(simp)
  4570 apply(simp)
  4578    apply(simp)
  4571 apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
  4579   apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
  4572 done
  4580   done
  4573 
  4581 
  4574 lemma psubst_nsubst:
  4582 lemma psubst_nsubst:
  4575   assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
  4583   assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
  4576   shows "((x,a,P)#\<theta>_n),\<theta>_c<M> = ((\<theta>_n,\<theta>_c<M>){x:=<a>.P})"
  4584   shows "((x,a,P)#\<theta>_n),\<theta>_c<M> = ((\<theta>_n,\<theta>_c<M>){x:=<a>.P})"
  4577 using a
  4585   using a
  4578 apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
  4586   apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
  4579 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4587              apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4580 apply(simp add: lookup_fresh)
  4588               apply(simp add: lookup_fresh)
  4581 apply(rule lookupb_lookupa)
  4589               apply(rule lookupb_lookupa)
  4582 apply(simp)
  4590               apply(simp)
  4583 apply(rule sym)
  4591              apply(rule sym)
  4584 apply(rule forget)
  4592              apply(rule forget)
  4585 apply(rule lookup_freshness)
  4593              apply(rule lookup_freshness)
  4586 apply(simp add: fresh_atm)
  4594              apply(simp add: fresh_atm)
  4587 apply(auto simp add: lookupc_freshness fresh_list_cons fresh_prod)[1]
  4595             apply(auto simp add: lookupc_freshness fresh_list_cons fresh_prod)[1]
  4588 apply(simp add: lookupc_fresh)
  4596                  apply(simp add: lookupc_fresh)
  4589 apply(rule sym)
  4597                  apply(rule sym)
  4590 apply(rule trans)
  4598                  apply(rule trans)
  4591 apply(rule better_Cut_substn)
  4599                   apply(rule better_Cut_substn)
  4592 apply(simp add: abs_fresh)
  4600                    apply(simp add: abs_fresh)
  4593 apply(simp add: abs_fresh fresh_atm)
  4601                   apply(simp add: abs_fresh fresh_atm)
  4594 apply(simp add: lookupd_fresh)
  4602                  apply(simp add: lookupd_fresh)
  4595 apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
  4603                  apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
  4596 apply(simp add: forget)
  4604                   apply(simp add: forget)
  4597 apply(simp add: trm.inject)
  4605                   apply(simp add: trm.inject)
  4598 apply(rule sym)
  4606                   apply(rule sym)
  4599 apply(simp add: alpha crename_swap fresh_atm)
  4607                   apply(simp add: alpha crename_swap fresh_atm)
  4600 apply(rule lookupd_freshness)
  4608                  apply(rule lookupd_freshness)
  4601 apply(simp add: fresh_atm)
  4609                  apply(simp add: fresh_atm)
  4602 apply(rule sym)
  4610                 apply(rule sym)
  4603 apply(rule trans)
  4611                 apply(rule trans)
  4604 apply(rule better_Cut_substn)
  4612                  apply(rule better_Cut_substn)
  4605 apply(simp add: abs_fresh) 
  4613                   apply(simp add: abs_fresh) 
  4606 apply(simp add: abs_fresh fresh_atm)
  4614                  apply(simp add: abs_fresh fresh_atm)
  4607 apply(simp)
  4615                 apply(simp)
  4608 apply(rule conjI)
  4616                 apply(rule conjI)
  4609 apply(rule impI)
  4617                  apply(rule impI)
  4610 apply(simp add: lookupc_unicity)
  4618                  apply(simp add: lookupc_unicity)
  4611 apply(rule impI)
  4619                 apply(rule impI)
  4612 apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
  4620                 apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
  4613 apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
  4621                  apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
  4614 apply(simp add: forget)
  4622                   apply(simp add: forget)
  4615 apply(rule lookupd_freshness)
  4623                  apply(rule lookupd_freshness)
  4616 apply(simp add: fresh_atm)
  4624                  apply(simp add: fresh_atm)
  4617 apply(rule lookupc_freshness)
  4625                 apply(rule lookupc_freshness)
  4618 apply(simp add: fresh_atm)
  4626                 apply(simp add: fresh_atm)
  4619 apply(rule sym)
  4627                apply(rule sym)
  4620 apply(rule trans)
  4628                apply(rule trans)
  4621 apply(rule better_Cut_substn)
  4629                 apply(rule better_Cut_substn)
  4622 apply(simp add: abs_fresh)
  4630                  apply(simp add: abs_fresh)
  4623 apply(simp add: abs_fresh fresh_atm)
  4631                 apply(simp add: abs_fresh fresh_atm)
  4624 apply(simp)
  4632                apply(simp)
  4625 apply(rule conjI)
  4633                apply(rule conjI)
  4626 apply(rule impI)
  4634                 apply(rule impI)
  4627 apply(simp add: trm.inject)
  4635                 apply(simp add: trm.inject)
  4628 apply(rule sym)
  4636                 apply(rule sym)
  4629 apply(simp add: alpha crename_swap fresh_atm)
  4637                 apply(simp add: alpha crename_swap fresh_atm)
  4630 apply(rule impI)
  4638                apply(rule impI)
  4631 apply(simp add: lookupc_fresh)
  4639                apply(simp add: lookupc_fresh)
  4632 apply(rule sym)
  4640               apply(rule sym)
  4633 apply(rule trans)
  4641               apply(rule trans)
  4634 apply(rule better_Cut_substn)
  4642                apply(rule better_Cut_substn)
  4635 apply(simp add: abs_fresh)
  4643                 apply(simp add: abs_fresh)
  4636 apply(simp add: abs_fresh fresh_atm)
  4644                apply(simp add: abs_fresh fresh_atm)
  4637 apply(simp)
  4645               apply(simp)
  4638 apply(rule conjI)
  4646               apply(rule conjI)
  4639 apply(rule impI)
  4647                apply(rule impI)
  4640 apply(simp add: lookupc_unicity)
  4648                apply(simp add: lookupc_unicity)
  4641 apply(rule impI)
  4649               apply(rule impI)
  4642 apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
  4650               apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
  4643 apply(simp add: forget)
  4651                apply(simp add: forget)
  4644 apply(rule lookupc_freshness)
  4652               apply(rule lookupc_freshness)
  4645 apply(simp add: fresh_prod fresh_atm)
  4653               apply(simp add: fresh_prod fresh_atm)
  4646 apply(rule sym)
  4654              apply(rule sym)
  4647 apply(rule trans)
  4655              apply(rule trans)
  4648 apply(rule better_Cut_substn)
  4656               apply(rule better_Cut_substn)
  4649 apply(simp add: abs_fresh)
  4657                apply(simp add: abs_fresh)
  4650 apply(simp add: abs_fresh fresh_atm)
  4658               apply(simp add: abs_fresh fresh_atm)
  4651 apply(simp)
  4659              apply(simp)
  4652 apply(rule conjI)
  4660              apply(rule conjI)
  4653 apply(rule impI)
  4661               apply(rule impI)
  4654 apply(drule ax_psubst)
  4662               apply(drule ax_psubst)
  4655 apply(simp)
  4663                 apply(simp)
  4656 apply(simp)
  4664                apply(simp)
  4657 apply(simp)
  4665               apply(simp)
  4658 apply(blast)
  4666               apply(blast)
  4659 apply(rule impI)
  4667              apply(rule impI)
  4660 apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
  4668              apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
  4661 apply(simp add: forget)
  4669               apply(simp add: forget)
  4662 apply(rule lookupd_freshness)
  4670              apply(rule lookupd_freshness)
  4663 apply(simp add: fresh_atm)
  4671              apply(simp add: fresh_atm)
  4664 apply(rule sym)
  4672             apply(rule sym)
  4665 apply(rule trans)
  4673             apply(rule trans)
  4666 apply(rule better_Cut_substn)
  4674              apply(rule better_Cut_substn)
  4667 apply(simp add: abs_fresh)
  4675               apply(simp add: abs_fresh)
  4668 apply(simp add: abs_fresh fresh_atm)
  4676              apply(simp add: abs_fresh fresh_atm)
  4669 apply(simp)
  4677             apply(simp)
  4670 apply(rule impI)
  4678             apply(rule impI)
  4671 apply(drule ax_psubst)
  4679             apply(drule ax_psubst)
  4672 apply(simp)
  4680               apply(simp)
  4673 apply(simp)
  4681              apply(simp)
  4674 apply(blast)
  4682             apply(blast)
  4675 (* NotR *)
  4683     (* NotR *)
  4676 apply(simp)
  4684            apply(simp)
  4677 apply(case_tac "findc \<theta>_c coname")
  4685            apply(case_tac "findc \<theta>_c coname")
  4678 apply(simp)
  4686             apply(simp)
  4679 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4687             apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4680 apply(simp)
  4688            apply(simp)
  4681 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4689            apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4682 apply(generate_fresh "coname")
  4690            apply(generate_fresh "coname")
  4683 apply(fresh_fun_simp)
  4691            apply(fresh_fun_simp)
  4684 apply(fresh_fun_simp)
  4692            apply(fresh_fun_simp)
  4685 apply(rule sym)
  4693            apply(rule sym)
  4686 apply(rule trans)
  4694            apply(rule trans)
  4687 apply(rule better_Cut_substn1)
  4695             apply(rule better_Cut_substn1)
  4688 apply(simp add: cmaps_fresh)
  4696               apply(simp add: cmaps_fresh)
  4689 apply(simp)
  4697              apply(simp)
  4690 apply(simp)
  4698             apply(simp)
  4691 apply(simp)
  4699            apply(simp)
  4692 (* NotL *)
  4700     (* NotL *)
  4693 apply(simp)
  4701           apply(simp)
  4694 apply(case_tac "findn \<theta>_n name")
  4702           apply(case_tac "findn \<theta>_n name")
  4695 apply(simp)
  4703            apply(simp)
  4696 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4704            apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4697 apply(simp)
  4705           apply(simp)
  4698 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4706           apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4699 apply(drule nmaps_false)
  4707            apply(drule nmaps_false)
  4700 apply(simp)
  4708             apply(simp)
  4701 apply(simp)
  4709            apply(simp)
  4702 apply(generate_fresh "name")
  4710           apply(generate_fresh "name")
  4703 apply(fresh_fun_simp)
  4711           apply(fresh_fun_simp)
  4704 apply(fresh_fun_simp)
  4712           apply(fresh_fun_simp)
  4705 apply(rule sym)
  4713           apply(rule sym)
  4706 apply(rule trans)
  4714           apply(rule trans)
  4707 apply(rule better_Cut_substn2)
  4715            apply(rule better_Cut_substn2)
  4708 apply(simp)
  4716             apply(simp)
  4709 apply(simp add: nmaps_fresh)
  4717            apply(simp add: nmaps_fresh)
  4710 apply(simp add: fresh_prod fresh_atm)
  4718           apply(simp add: fresh_prod fresh_atm)
  4711 (* AndR *)
  4719     (* AndR *)
  4712 apply(simp)
  4720          apply(simp)
  4713 apply(case_tac "findc \<theta>_c coname3")
  4721          apply(case_tac "findc \<theta>_c coname3")
  4714 apply(simp)
  4722           apply(simp)
  4715 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4723           apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4716 apply(simp)
  4724          apply(simp)
  4717 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4725          apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4718 apply(generate_fresh "coname")
  4726          apply(generate_fresh "coname")
  4719 apply(fresh_fun_simp)
  4727          apply(fresh_fun_simp)
  4720 apply(fresh_fun_simp)
  4728          apply(fresh_fun_simp)
  4721 apply(rule sym)
  4729          apply(rule sym)
  4722 apply(rule trans)
  4730          apply(rule trans)
  4723 apply(rule better_Cut_substn1)
  4731           apply(rule better_Cut_substn1)
  4724 apply(simp add: cmaps_fresh)
  4732             apply(simp add: cmaps_fresh)
  4725 apply(simp)
  4733            apply(simp)
  4726 apply(simp)
  4734           apply(simp)
  4727 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4735          apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4728 (* AndL1 *)
  4736     (* AndL1 *)
  4729 apply(simp)
  4737         apply(simp)
  4730 apply(case_tac "findn \<theta>_n name2")
  4738         apply(case_tac "findn \<theta>_n name2")
  4731 apply(simp)
  4739          apply(simp)
  4732 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4740          apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4733 apply(simp)
  4741         apply(simp)
  4734 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4742         apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4735 apply(drule nmaps_false)
  4743          apply(drule nmaps_false)
  4736 apply(simp)
  4744           apply(simp)
  4737 apply(simp)
  4745          apply(simp)
  4738 apply(generate_fresh "name")
  4746         apply(generate_fresh "name")
  4739 apply(fresh_fun_simp)
  4747         apply(fresh_fun_simp)
  4740 apply(fresh_fun_simp)
  4748         apply(fresh_fun_simp)
  4741 apply(rule sym)
  4749         apply(rule sym)
  4742 apply(rule trans)
  4750         apply(rule trans)
  4743 apply(rule better_Cut_substn2)
  4751          apply(rule better_Cut_substn2)
  4744 apply(simp)
  4752           apply(simp)
  4745 apply(simp add: nmaps_fresh)
  4753          apply(simp add: nmaps_fresh)
  4746 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4754         apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4747 (* AndL2 *)
  4755     (* AndL2 *)
  4748 apply(simp)
  4756        apply(simp)
  4749 apply(case_tac "findn \<theta>_n name2")
  4757        apply(case_tac "findn \<theta>_n name2")
  4750 apply(simp)
  4758         apply(simp)
  4751 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4759         apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4752 apply(simp)
  4760        apply(simp)
  4753 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4761        apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4754 apply(drule nmaps_false)
  4762         apply(drule nmaps_false)
  4755 apply(simp)
  4763          apply(simp)
  4756 apply(simp)
  4764         apply(simp)
  4757 apply(generate_fresh "name")
  4765        apply(generate_fresh "name")
  4758 apply(fresh_fun_simp)
  4766        apply(fresh_fun_simp)
  4759 apply(fresh_fun_simp)
  4767        apply(fresh_fun_simp)
  4760 apply(rule sym)
  4768        apply(rule sym)
  4761 apply(rule trans)
  4769        apply(rule trans)
  4762 apply(rule better_Cut_substn2)
  4770         apply(rule better_Cut_substn2)
  4763 apply(simp)
  4771          apply(simp)
  4764 apply(simp add: nmaps_fresh)
  4772         apply(simp add: nmaps_fresh)
  4765 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4773        apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4766 (* OrR1 *)
  4774     (* OrR1 *)
  4767 apply(simp)
  4775       apply(simp)
  4768 apply(case_tac "findc \<theta>_c coname2")
  4776       apply(case_tac "findc \<theta>_c coname2")
  4769 apply(simp)
  4777        apply(simp)
  4770 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4778        apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4771 apply(simp)
  4779       apply(simp)
  4772 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4780       apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4773 apply(generate_fresh "coname")
  4781       apply(generate_fresh "coname")
  4774 apply(fresh_fun_simp)
  4782       apply(fresh_fun_simp)
  4775 apply(fresh_fun_simp)
  4783       apply(fresh_fun_simp)
  4776 apply(rule sym)
  4784       apply(rule sym)
  4777 apply(rule trans)
  4785       apply(rule trans)
  4778 apply(rule better_Cut_substn1)
  4786        apply(rule better_Cut_substn1)
  4779 apply(simp add: cmaps_fresh)
  4787          apply(simp add: cmaps_fresh)
  4780 apply(simp)
  4788         apply(simp)
  4781 apply(simp)
  4789        apply(simp)
  4782 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4790       apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4783 (* OrR2 *)
  4791     (* OrR2 *)
  4784 apply(simp)
  4792      apply(simp)
  4785 apply(case_tac "findc \<theta>_c coname2")
  4793      apply(case_tac "findc \<theta>_c coname2")
  4786 apply(simp)
  4794       apply(simp)
  4787 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4795       apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4788 apply(simp)
  4796      apply(simp)
  4789 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4797      apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4790 apply(generate_fresh "coname")
  4798      apply(generate_fresh "coname")
  4791 apply(fresh_fun_simp)
  4799      apply(fresh_fun_simp)
  4792 apply(fresh_fun_simp)
  4800      apply(fresh_fun_simp)
  4793 apply(rule sym)
  4801      apply(rule sym)
  4794 apply(rule trans)
  4802      apply(rule trans)
  4795 apply(rule better_Cut_substn1)
  4803       apply(rule better_Cut_substn1)
  4796 apply(simp add: cmaps_fresh)
  4804         apply(simp add: cmaps_fresh)
  4797 apply(simp)
  4805        apply(simp)
  4798 apply(simp)
  4806       apply(simp)
  4799 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4807      apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4800 (* OrL *)
  4808     (* OrL *)
  4801 apply(simp)
  4809     apply(simp)
  4802 apply(case_tac "findn \<theta>_n name3")
  4810     apply(case_tac "findn \<theta>_n name3")
  4803 apply(simp)
  4811      apply(simp)
  4804 apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
  4812      apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
  4805 apply(simp)
  4813     apply(simp)
  4806 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4814     apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4807 apply(drule nmaps_false)
  4815      apply(drule nmaps_false)
  4808 apply(simp)
  4816       apply(simp)
  4809 apply(simp)
  4817      apply(simp)
  4810 apply(generate_fresh "name")
  4818     apply(generate_fresh "name")
  4811 apply(fresh_fun_simp)
  4819     apply(fresh_fun_simp)
  4812 apply(fresh_fun_simp)
  4820     apply(fresh_fun_simp)
  4813 apply(rule sym)
  4821     apply(rule sym)
  4814 apply(rule trans)
  4822     apply(rule trans)
  4815 apply(rule better_Cut_substn2)
  4823      apply(rule better_Cut_substn2)
  4816 apply(simp)
  4824       apply(simp)
  4817 apply(simp add: nmaps_fresh)
  4825      apply(simp add: nmaps_fresh)
  4818 apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
  4826     apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
  4819 (* ImpR *)
  4827     (* ImpR *)
  4820 apply(simp)
  4828    apply(simp)
  4821 apply(case_tac "findc \<theta>_c coname2")
  4829    apply(case_tac "findc \<theta>_c coname2")
  4822 apply(simp)
  4830     apply(simp)
  4823 apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4831     apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
  4824 apply(simp)
  4832    apply(simp)
  4825 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4833    apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4826 apply(generate_fresh "coname")
  4834    apply(generate_fresh "coname")
  4827 apply(fresh_fun_simp)
  4835    apply(fresh_fun_simp)
  4828 apply(fresh_fun_simp)
  4836    apply(fresh_fun_simp)
  4829 apply(rule sym)
  4837    apply(rule sym)
  4830 apply(rule trans)
  4838    apply(rule trans)
  4831 apply(rule better_Cut_substn1)
  4839     apply(rule better_Cut_substn1)
  4832 apply(simp add: cmaps_fresh)
  4840       apply(simp add: cmaps_fresh)
  4833 apply(simp)
  4841      apply(simp)
  4834 apply(simp)
  4842     apply(simp)
  4835 apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4843    apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
  4836 (* ImpL *)
  4844     (* ImpL *)
  4837 apply(simp)
  4845   apply(simp)
  4838 apply(case_tac "findn \<theta>_n name2")
  4846   apply(case_tac "findn \<theta>_n name2")
  4839 apply(simp)
  4847    apply(simp)
  4840 apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
  4848    apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
  4841 apply(simp)
  4849   apply(simp)
  4842 apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4850   apply(auto simp add: fresh_list_cons fresh_prod)[1]
  4843 apply(drule nmaps_false)
  4851    apply(drule nmaps_false)
  4844 apply(simp)
  4852     apply(simp)
  4845 apply(simp)
  4853    apply(simp)
  4846 apply(generate_fresh "name")
  4854   apply(generate_fresh "name")
  4847 apply(fresh_fun_simp)
  4855   apply(fresh_fun_simp)
  4848 apply(fresh_fun_simp)
  4856   apply(fresh_fun_simp)
  4849 apply(rule sym)
  4857   apply(rule sym)
  4850 apply(rule trans)
  4858   apply(rule trans)
  4851 apply(rule better_Cut_substn2)
  4859    apply(rule better_Cut_substn2)
  4852 apply(simp)
  4860     apply(simp)
  4853 apply(simp add: nmaps_fresh)
  4861    apply(simp add: nmaps_fresh)
  4854 apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
  4862   apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
  4855 done
  4863   done
  4856 
  4864 
  4857 definition 
  4865 definition 
  4858   ncloses :: "(name\<times>coname\<times>trm) list\<Rightarrow>(name\<times>ty) list \<Rightarrow> bool" ("_ ncloses _" [55,55] 55) 
  4866   ncloses :: "(name\<times>coname\<times>trm) list\<Rightarrow>(name\<times>ty) list \<Rightarrow> bool" ("_ ncloses _" [55,55] 55) 
  4859 where
  4867   where
  4860   "\<theta>_n ncloses \<Gamma> \<equiv> \<forall>x B. ((x,B) \<in> set \<Gamma> \<longrightarrow> (\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)))"
  4868     "\<theta>_n ncloses \<Gamma> \<equiv> \<forall>x B. ((x,B) \<in> set \<Gamma> \<longrightarrow> (\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)))"
  4861   
  4869 
  4862 definition 
  4870 definition 
  4863   ccloses :: "(coname\<times>name\<times>trm) list\<Rightarrow>(coname\<times>ty) list \<Rightarrow> bool" ("_ ccloses _" [55,55] 55) 
  4871   ccloses :: "(coname\<times>name\<times>trm) list\<Rightarrow>(coname\<times>ty) list \<Rightarrow> bool" ("_ ccloses _" [55,55] 55) 
  4864 where
  4872   where
  4865   "\<theta>_c ccloses \<Delta> \<equiv> \<forall>a B. ((a,B) \<in> set \<Delta> \<longrightarrow> (\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)))"
  4873     "\<theta>_c ccloses \<Delta> \<equiv> \<forall>a B. ((a,B) \<in> set \<Delta> \<longrightarrow> (\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)))"
  4866 
  4874 
  4867 lemma ncloses_elim:
  4875 lemma ncloses_elim:
  4868   assumes a: "(x,B) \<in> set \<Gamma>"
  4876   assumes a: "(x,B) \<in> set \<Gamma>"
  4869   and     b: "\<theta>_n ncloses \<Gamma>"
  4877     and     b: "\<theta>_n ncloses \<Gamma>"
  4870   shows "\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)"
  4878   shows "\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)"
  4871 using a b by (auto simp add: ncloses_def)
  4879   using a b by (auto simp add: ncloses_def)
  4872 
  4880 
  4873 lemma ccloses_elim:
  4881 lemma ccloses_elim:
  4874   assumes a: "(a,B) \<in> set \<Delta>"
  4882   assumes a: "(a,B) \<in> set \<Delta>"
  4875   and     b: "\<theta>_c ccloses \<Delta>"
  4883     and     b: "\<theta>_c ccloses \<Delta>"
  4876   shows "\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)"
  4884   shows "\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)"
  4877 using a b by (auto simp add: ccloses_def)
  4885   using a b by (auto simp add: ccloses_def)
  4878 
  4886 
  4879 lemma ncloses_subset:
  4887 lemma ncloses_subset:
  4880   assumes a: "\<theta>_n ncloses \<Gamma>"
  4888   assumes a: "\<theta>_n ncloses \<Gamma>"
  4881   and     b: "set \<Gamma>' \<subseteq> set \<Gamma>"
  4889     and     b: "set \<Gamma>' \<subseteq> set \<Gamma>"
  4882   shows "\<theta>_n ncloses \<Gamma>'"
  4890   shows "\<theta>_n ncloses \<Gamma>'"
  4883 using a b by (auto  simp add: ncloses_def)
  4891   using a b by (auto  simp add: ncloses_def)
  4884 
  4892 
  4885 lemma ccloses_subset:
  4893 lemma ccloses_subset:
  4886   assumes a: "\<theta>_c ccloses \<Delta>"
  4894   assumes a: "\<theta>_c ccloses \<Delta>"
  4887   and     b: "set \<Delta>' \<subseteq> set \<Delta>"
  4895     and     b: "set \<Delta>' \<subseteq> set \<Delta>"
  4888   shows "\<theta>_c ccloses \<Delta>'"
  4896   shows "\<theta>_c ccloses \<Delta>'"
  4889 using a b by (auto  simp add: ccloses_def)
  4897   using a b by (auto  simp add: ccloses_def)
  4890 
  4898 
  4891 lemma validc_fresh:
  4899 lemma validc_fresh:
  4892   fixes a::"coname"
  4900   fixes a::"coname"
  4893   and   \<Delta>::"(coname\<times>ty) list"
  4901     and   \<Delta>::"(coname\<times>ty) list"
  4894   assumes a: "a\<sharp>\<Delta>"
  4902   assumes a: "a\<sharp>\<Delta>"
  4895   shows "\<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
  4903   shows "\<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
  4896 using a
  4904   using a
  4897 apply(induct \<Delta>)
  4905   apply(induct \<Delta>)
  4898 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  4906    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  4899 done
  4907   done
  4900 
  4908 
  4901 lemma validn_fresh:
  4909 lemma validn_fresh:
  4902   fixes x::"name"
  4910   fixes x::"name"
  4903   and   \<Gamma>::"(name\<times>ty) list"
  4911     and   \<Gamma>::"(name\<times>ty) list"
  4904   assumes a: "x\<sharp>\<Gamma>"
  4912   assumes a: "x\<sharp>\<Gamma>"
  4905   shows "\<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
  4913   shows "\<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
  4906 using a
  4914   using a
  4907 apply(induct \<Gamma>)
  4915   apply(induct \<Gamma>)
  4908 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  4916    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  4909 done
  4917   done
  4910 
  4918 
  4911 lemma ccloses_extend:
  4919 lemma ccloses_extend:
  4912   assumes a: "\<theta>_c ccloses \<Delta>" "a\<sharp>\<Delta>" "a\<sharp>\<theta>_c" "(x):P\<in>\<parallel>(B)\<parallel>"
  4920   assumes a: "\<theta>_c ccloses \<Delta>" "a\<sharp>\<Delta>" "a\<sharp>\<theta>_c" "(x):P\<in>\<parallel>(B)\<parallel>"
  4913   shows "(a,x,P)#\<theta>_c ccloses (a,B)#\<Delta>"
  4921   shows "(a,x,P)#\<theta>_c ccloses (a,B)#\<Delta>"
  4914 using a
  4922   using a
  4915 apply(simp add: ccloses_def)
  4923   apply(simp add: ccloses_def)
  4916 apply(drule validc_fresh)
  4924   apply(drule validc_fresh)
  4917 apply(auto)
  4925   apply(auto)
  4918 done
  4926   done
  4919 
  4927 
  4920 lemma ncloses_extend:
  4928 lemma ncloses_extend:
  4921   assumes a: "\<theta>_n ncloses \<Gamma>" "x\<sharp>\<Gamma>" "x\<sharp>\<theta>_n" "<a>:P\<in>\<parallel><B>\<parallel>"
  4929   assumes a: "\<theta>_n ncloses \<Gamma>" "x\<sharp>\<Gamma>" "x\<sharp>\<theta>_n" "<a>:P\<in>\<parallel><B>\<parallel>"
  4922   shows "(x,a,P)#\<theta>_n ncloses (x,B)#\<Gamma>"
  4930   shows "(x,a,P)#\<theta>_n ncloses (x,B)#\<Gamma>"
  4923 using a
  4931   using a
  4924 apply(simp add: ncloses_def)
  4932   apply(simp add: ncloses_def)
  4925 apply(drule validn_fresh)
  4933   apply(drule validn_fresh)
  4926 apply(auto)
  4934   apply(auto)
  4927 done
  4935   done
  4928 
  4936 
  4929 
  4937 
  4930 text \<open>typing relation\<close>
  4938 text \<open>typing relation\<close>
  4931 
  4939 
  4932 inductive
  4940 inductive
  4933    typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
  4941   typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
  4934 where
  4942   where
  4935   TAx:    "\<lbrakk>validn \<Gamma>;validc \<Delta>; (x,B)\<in>set \<Gamma>; (a,B)\<in>set \<Delta>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Ax x a \<turnstile> \<Delta>"
  4943     TAx:    "\<lbrakk>validn \<Gamma>;validc \<Delta>; (x,B)\<in>set \<Gamma>; (a,B)\<in>set \<Delta>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Ax x a \<turnstile> \<Delta>"
  4936 | TNotR:  "\<lbrakk>x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Delta>' = {(a,NOT B)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4944   | TNotR:  "\<lbrakk>x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Delta>' = {(a,NOT B)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4937            \<Longrightarrow> \<Gamma> \<turnstile> NotR (x).M a \<turnstile> \<Delta>'"
  4945            \<Longrightarrow> \<Gamma> \<turnstile> NotR (x).M a \<turnstile> \<Delta>'"
  4938 | TNotL:  "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); set \<Gamma>' = {(x,NOT B)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk>  
  4946   | TNotL:  "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); set \<Gamma>' = {(x,NOT B)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk>  
  4939            \<Longrightarrow> \<Gamma>' \<turnstile> NotL <a>.M x \<turnstile> \<Delta>"
  4947            \<Longrightarrow> \<Gamma>' \<turnstile> NotL <a>.M x \<turnstile> \<Delta>"
  4940 | TAndL1: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B1)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4948   | TAndL1: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B1)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4941            \<Longrightarrow> \<Gamma>' \<turnstile> AndL1 (x).M y \<turnstile> \<Delta>"
  4949            \<Longrightarrow> \<Gamma>' \<turnstile> AndL1 (x).M y \<turnstile> \<Delta>"
  4942 | TAndL2: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B2)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4950   | TAndL2: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B2)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4943            \<Longrightarrow> \<Gamma>' \<turnstile> AndL2 (x).M y \<turnstile> \<Delta>"
  4951            \<Longrightarrow> \<Gamma>' \<turnstile> AndL2 (x).M y \<turnstile> \<Delta>"
  4944 | TAndR:  "\<lbrakk>a\<sharp>(\<Delta>,N,c); b\<sharp>(\<Delta>,M,c); a\<noteq>b; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); \<Gamma> \<turnstile> N \<turnstile> ((b,C)#\<Delta>); 
  4952   | TAndR:  "\<lbrakk>a\<sharp>(\<Delta>,N,c); b\<sharp>(\<Delta>,M,c); a\<noteq>b; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); \<Gamma> \<turnstile> N \<turnstile> ((b,C)#\<Delta>); 
  4945            set \<Delta>' = {(c,B AND C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4953            set \<Delta>' = {(c,B AND C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4946            \<Longrightarrow> \<Gamma> \<turnstile> AndR <a>.M <b>.N c \<turnstile> \<Delta>'"
  4954            \<Longrightarrow> \<Gamma> \<turnstile> AndR <a>.M <b>.N c \<turnstile> \<Delta>'"
  4947 | TOrL:   "\<lbrakk>x\<sharp>(\<Gamma>,N,z); y\<sharp>(\<Gamma>,M,z); x\<noteq>y; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; ((y,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
  4955   | TOrL:   "\<lbrakk>x\<sharp>(\<Gamma>,N,z); y\<sharp>(\<Gamma>,M,z); x\<noteq>y; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; ((y,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
  4948            set \<Gamma>' = {(z,B OR C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4956            set \<Gamma>' = {(z,B OR C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4949            \<Longrightarrow> \<Gamma>' \<turnstile> OrL (x).M (y).N z \<turnstile> \<Delta>"
  4957            \<Longrightarrow> \<Gamma>' \<turnstile> OrL (x).M (y).N z \<turnstile> \<Delta>"
  4950 | TOrR1:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B1)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4958   | TOrR1:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B1)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4951            \<Longrightarrow> \<Gamma> \<turnstile> OrR1 <a>.M b \<turnstile> \<Delta>'"
  4959            \<Longrightarrow> \<Gamma> \<turnstile> OrR1 <a>.M b \<turnstile> \<Delta>'"
  4952 | TOrR2:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B2)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4960   | TOrR2:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B2)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4953            \<Longrightarrow> \<Gamma> \<turnstile> OrR2 <a>.M b \<turnstile> \<Delta>'"
  4961            \<Longrightarrow> \<Gamma> \<turnstile> OrR2 <a>.M b \<turnstile> \<Delta>'"
  4954 | TImpL:  "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M,y); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
  4962   | TImpL:  "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M,y); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
  4955            set \<Gamma>' = {(y,B IMP C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4963            set \<Gamma>' = {(y,B IMP C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
  4956            \<Longrightarrow> \<Gamma>' \<turnstile> ImpL <a>.M (x).N y \<turnstile> \<Delta>"
  4964            \<Longrightarrow> \<Gamma>' \<turnstile> ImpL <a>.M (x).N y \<turnstile> \<Delta>"
  4957 | TImpR:  "\<lbrakk>a\<sharp>(\<Delta>,b); x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> ((a,C)#\<Delta>); set \<Delta>' = {(b,B IMP C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4965   | TImpR:  "\<lbrakk>a\<sharp>(\<Delta>,b); x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> ((a,C)#\<Delta>); set \<Delta>' = {(b,B IMP C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
  4958            \<Longrightarrow> \<Gamma> \<turnstile> ImpR (x).<a>.M b \<turnstile> \<Delta>'"
  4966            \<Longrightarrow> \<Gamma> \<turnstile> ImpR (x).<a>.M b \<turnstile> \<Delta>'"
  4959 | TCut:   "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,B)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>\<rbrakk> 
  4967   | TCut:   "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,B)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>\<rbrakk> 
  4960            \<Longrightarrow> \<Gamma> \<turnstile> Cut <a>.M (x).N \<turnstile> \<Delta>"
  4968            \<Longrightarrow> \<Gamma> \<turnstile> Cut <a>.M (x).N \<turnstile> \<Delta>"
  4961 
  4969 
  4962 equivariance typing
  4970 equivariance typing
  4963 
  4971 
  4964 lemma fresh_set_member:
  4972 lemma fresh_set_member:
  4965   fixes x::"name"
  4973   fixes x::"name"
  4966   and   a::"coname"
  4974     and   a::"coname"
  4967   shows "x\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> x\<sharp>e"
  4975   shows "x\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> x\<sharp>e"
  4968   and   "a\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> a\<sharp>e"   
  4976     and   "a\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> a\<sharp>e"   
  4969 by (induct L) (auto simp add: fresh_list_cons) 
  4977   by (induct L) (auto simp add: fresh_list_cons) 
  4970 
  4978 
  4971 lemma fresh_subset:
  4979 lemma fresh_subset:
  4972   fixes x::"name"
  4980   fixes x::"name"
  4973   and   a::"coname"
  4981     and   a::"coname"
  4974   shows "x\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> x\<sharp>L'"
  4982   shows "x\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> x\<sharp>L'"
  4975   and   "a\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> a\<sharp>L'"   
  4983     and   "a\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> a\<sharp>L'"   
  4976 apply(induct L' arbitrary: L) 
  4984    apply(induct L' arbitrary: L) 
  4977 apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
  4985      apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
  4978 done
  4986   done
  4979 
  4987 
  4980 lemma fresh_subset_ext:
  4988 lemma fresh_subset_ext:
  4981   fixes x::"name"
  4989   fixes x::"name"
  4982   and   a::"coname"
  4990     and   a::"coname"
  4983   shows "x\<sharp>L \<Longrightarrow> x\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> x\<sharp>L'"
  4991   shows "x\<sharp>L \<Longrightarrow> x\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> x\<sharp>L'"
  4984   and   "a\<sharp>L \<Longrightarrow> a\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> a\<sharp>L'"   
  4992     and   "a\<sharp>L \<Longrightarrow> a\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> a\<sharp>L'"   
  4985 apply(induct L' arbitrary: L) 
  4993    apply(induct L' arbitrary: L) 
  4986 apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
  4994      apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
  4987 done
  4995   done
  4988 
  4996 
  4989 lemma fresh_under_insert:
  4997 lemma fresh_under_insert:
  4990   fixes x::"name"
  4998   fixes x::"name"
  4991   and   a::"coname"
  4999     and   a::"coname"
  4992   and   \<Gamma>::"ctxtn"
  5000     and   \<Gamma>::"ctxtn"
  4993   and   \<Delta>::"ctxtc"
  5001     and   \<Delta>::"ctxtc"
  4994   shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<noteq>y \<Longrightarrow> set \<Gamma>' = insert (y,B) (set \<Gamma>) \<Longrightarrow> x\<sharp>\<Gamma>'"
  5002   shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<noteq>y \<Longrightarrow> set \<Gamma>' = insert (y,B) (set \<Gamma>) \<Longrightarrow> x\<sharp>\<Gamma>'"
  4995   and   "a\<sharp>\<Delta> \<Longrightarrow> a\<noteq>c \<Longrightarrow> set \<Delta>' = insert (c,B) (set \<Delta>) \<Longrightarrow> a\<sharp>\<Delta>'"   
  5003     and   "a\<sharp>\<Delta> \<Longrightarrow> a\<noteq>c \<Longrightarrow> set \<Delta>' = insert (c,B) (set \<Delta>) \<Longrightarrow> a\<sharp>\<Delta>'"   
  4996 apply(rule fresh_subset_ext(1))
  5004    apply(rule fresh_subset_ext(1))
  4997 apply(auto simp add: fresh_prod fresh_atm fresh_ty)
  5005      apply(auto simp add: fresh_prod fresh_atm fresh_ty)
  4998 apply(rule fresh_subset_ext(2))
  5006   apply(rule fresh_subset_ext(2))
  4999 apply(auto simp add: fresh_prod fresh_atm fresh_ty)
  5007     apply(auto simp add: fresh_prod fresh_atm fresh_ty)
  5000 done
  5008   done
  5001 
  5009 
  5002 nominal_inductive typing
  5010 nominal_inductive typing
  5003   apply (simp_all add: abs_fresh fresh_atm fresh_list_cons fresh_prod fresh_ty fresh_ctxt 
  5011                        apply (simp_all add: abs_fresh fresh_atm fresh_list_cons fresh_prod fresh_ty fresh_ctxt 
  5004                        fresh_list_append abs_supp fin_supp)
  5012       fresh_list_append abs_supp fin_supp)
  5005   apply(auto intro: fresh_under_insert)
  5013            apply(auto intro: fresh_under_insert)
  5006   done
  5014   done
  5007 
  5015 
  5008 lemma validn_elim:
  5016 lemma validn_elim:
  5009   assumes a: "validn ((x,B)#\<Gamma>)"
  5017   assumes a: "validn ((x,B)#\<Gamma>)"
  5010   shows "validn \<Gamma> \<and> x\<sharp>\<Gamma>"
  5018   shows "validn \<Gamma> \<and> x\<sharp>\<Gamma>"
  5011 using a
  5019   using a
  5012 apply(erule_tac validn.cases)
  5020   apply(erule_tac validn.cases)
  5013 apply(auto)
  5021    apply(auto)
  5014 done
  5022   done
  5015 
  5023 
  5016 lemma validc_elim:
  5024 lemma validc_elim:
  5017   assumes a: "validc ((a,B)#\<Delta>)"
  5025   assumes a: "validc ((a,B)#\<Delta>)"
  5018   shows "validc \<Delta> \<and> a\<sharp>\<Delta>"
  5026   shows "validc \<Delta> \<and> a\<sharp>\<Delta>"
  5019 using a
  5027   using a
  5020 apply(erule_tac validc.cases)
  5028   apply(erule_tac validc.cases)
  5021 apply(auto)
  5029    apply(auto)
  5022 done
  5030   done
  5023 
  5031 
  5024 lemma context_fresh:
  5032 lemma context_fresh:
  5025   fixes x::"name"
  5033   fixes x::"name"
  5026   and   a::"coname"
  5034     and   a::"coname"
  5027   shows "x\<sharp>\<Gamma> \<Longrightarrow> \<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
  5035   shows "x\<sharp>\<Gamma> \<Longrightarrow> \<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
  5028   and   "a\<sharp>\<Delta> \<Longrightarrow> \<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
  5036     and   "a\<sharp>\<Delta> \<Longrightarrow> \<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
  5029 apply -
  5037    apply -
  5030 apply(induct \<Gamma>)
  5038    apply(induct \<Gamma>)
  5031 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  5039     apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  5032 apply(induct \<Delta>)
  5040   apply(induct \<Delta>)
  5033 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  5041    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  5034 done
  5042   done
  5035 
  5043 
  5036 lemma typing_implies_valid:
  5044 lemma typing_implies_valid:
  5037   assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
  5045   assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
  5038   shows "validn \<Gamma> \<and> validc \<Delta>"
  5046   shows "validn \<Gamma> \<and> validc \<Delta>"
  5039 using a
  5047   using a
  5040 apply(nominal_induct rule: typing.strong_induct)
  5048   apply(nominal_induct rule: typing.strong_induct)
  5041 apply(auto dest: validn_elim validc_elim)
  5049              apply(auto dest: validn_elim validc_elim)
  5042 done
  5050   done
  5043 
  5051 
  5044 lemma ty_perm:
  5052 lemma ty_perm:
  5045   fixes pi1::"name prm"
  5053   fixes pi1::"name prm"
  5046   and   pi2::"coname prm"
  5054     and   pi2::"coname prm"
  5047   and   B::"ty"
  5055     and   B::"ty"
  5048   shows "pi1\<bullet>B=B" and "pi2\<bullet>B=B"
  5056   shows "pi1\<bullet>B=B" and "pi2\<bullet>B=B"
  5049 apply(nominal_induct B rule: ty.strong_induct)
  5057    apply(nominal_induct B rule: ty.strong_induct)
  5050 apply(auto simp add: perm_string)
  5058            apply(auto simp add: perm_string)
  5051 done
  5059   done
  5052 
  5060 
  5053 lemma ctxt_perm:
  5061 lemma ctxt_perm:
  5054   fixes pi1::"name prm"
  5062   fixes pi1::"name prm"
  5055   and   pi2::"coname prm"
  5063     and   pi2::"coname prm"
  5056   and   \<Gamma>::"ctxtn"
  5064     and   \<Gamma>::"ctxtn"
  5057   and   \<Delta>::"ctxtc"
  5065     and   \<Delta>::"ctxtc"
  5058   shows "pi2\<bullet>\<Gamma>=\<Gamma>" and "pi1\<bullet>\<Delta>=\<Delta>"
  5066   shows "pi2\<bullet>\<Gamma>=\<Gamma>" and "pi1\<bullet>\<Delta>=\<Delta>"
  5059 apply -
  5067    apply -
  5060 apply(induct \<Gamma>)
  5068    apply(induct \<Gamma>)
  5061 apply(auto simp add: calc_atm ty_perm)
  5069     apply(auto simp add: calc_atm ty_perm)
  5062 apply(induct \<Delta>)
  5070   apply(induct \<Delta>)
  5063 apply(auto simp add: calc_atm ty_perm)
  5071    apply(auto simp add: calc_atm ty_perm)
  5064 done
  5072   done
  5065 
  5073 
  5066 lemma typing_Ax_elim1: 
  5074 lemma typing_Ax_elim1: 
  5067   assumes a: "\<Gamma> \<turnstile> Ax x a \<turnstile> ((a,B)#\<Delta>)"
  5075   assumes a: "\<Gamma> \<turnstile> Ax x a \<turnstile> ((a,B)#\<Delta>)"
  5068   shows "(x,B)\<in>set \<Gamma>"
  5076   shows "(x,B)\<in>set \<Gamma>"
  5069 using a
  5077   using a
  5070 apply(erule_tac typing.cases)
  5078   apply(erule_tac typing.cases)
  5071 apply(simp_all add: trm.inject)
  5079              apply(simp_all add: trm.inject)
  5072 apply(auto)
  5080   apply(auto)
  5073 apply(auto dest: validc_elim context_fresh)
  5081   apply(auto dest: validc_elim context_fresh)
  5074 done
  5082   done
  5075 
  5083 
  5076 lemma typing_Ax_elim2: 
  5084 lemma typing_Ax_elim2: 
  5077   assumes a: "((x,B)#\<Gamma>) \<turnstile> Ax x a \<turnstile> \<Delta>"
  5085   assumes a: "((x,B)#\<Gamma>) \<turnstile> Ax x a \<turnstile> \<Delta>"
  5078   shows "(a,B)\<in>set \<Delta>"
  5086   shows "(a,B)\<in>set \<Delta>"
  5079 using a
  5087   using a
  5080 apply(erule_tac typing.cases)
  5088   apply(erule_tac typing.cases)
  5081 apply(simp_all add: trm.inject)
  5089              apply(simp_all add: trm.inject)
  5082 apply(auto  dest!: validn_elim context_fresh)
  5090   apply(auto  dest!: validn_elim context_fresh)
  5083 done
  5091   done
  5084 
  5092 
  5085 lemma psubst_Ax_aux: 
  5093 lemma psubst_Ax_aux: 
  5086   assumes a: "\<theta>_c cmaps a to Some (y,N)"
  5094   assumes a: "\<theta>_c cmaps a to Some (y,N)"
  5087   shows "lookupb x a \<theta>_c c P = Cut <c>.P (y).N"
  5095   shows "lookupb x a \<theta>_c c P = Cut <c>.P (y).N"
  5088 using a
  5096   using a
  5089 apply(induct \<theta>_c)
  5097   apply(induct \<theta>_c)
  5090 apply(auto)
  5098    apply(auto)
  5091 done
  5099   done
  5092 
  5100 
  5093 lemma psubst_Ax:
  5101 lemma psubst_Ax:
  5094   assumes a: "\<theta>_n nmaps x to Some (c,P)"
  5102   assumes a: "\<theta>_n nmaps x to Some (c,P)"
  5095   and     b: "\<theta>_c cmaps a to Some (y,N)"
  5103     and     b: "\<theta>_c cmaps a to Some (y,N)"
  5096   shows "\<theta>_n,\<theta>_c<Ax x a> = Cut <c>.P (y).N"
  5104   shows "\<theta>_n,\<theta>_c<Ax x a> = Cut <c>.P (y).N"
  5097 using a b
  5105   using a b
  5098 apply(induct \<theta>_n)
  5106   apply(induct \<theta>_n)
  5099 apply(auto simp add: psubst_Ax_aux)
  5107    apply(auto simp add: psubst_Ax_aux)
  5100 done
  5108   done
  5101 
  5109 
  5102 lemma psubst_Cut:
  5110 lemma psubst_Cut:
  5103   assumes a: "\<forall>x. M\<noteq>Ax x c"
  5111   assumes a: "\<forall>x. M\<noteq>Ax x c"
  5104   and     b: "\<forall>a. N\<noteq>Ax x a"
  5112     and     b: "\<forall>a. N\<noteq>Ax x a"
  5105   and     c: "c\<sharp>(\<theta>_n,\<theta>_c,N)" "x\<sharp>(\<theta>_n,\<theta>_c,M)"
  5113     and     c: "c\<sharp>(\<theta>_n,\<theta>_c,N)" "x\<sharp>(\<theta>_n,\<theta>_c,M)"
  5106   shows "\<theta>_n,\<theta>_c<Cut <c>.M (x).N> = Cut <c>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>)"
  5114   shows "\<theta>_n,\<theta>_c<Cut <c>.M (x).N> = Cut <c>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>)"
  5107 using a b c
  5115   using a b c
  5108 apply(simp)
  5116   apply(simp)
  5109 done
  5117   done
  5110 
  5118 
  5111 lemma all_CAND: 
  5119 lemma all_CAND: 
  5112   assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
  5120   assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
  5113   and     b: "\<theta>_n ncloses \<Gamma>"
  5121     and     b: "\<theta>_n ncloses \<Gamma>"
  5114   and     c: "\<theta>_c ccloses \<Delta>"
  5122     and     c: "\<theta>_c ccloses \<Delta>"
  5115   shows "SNa (\<theta>_n,\<theta>_c<M>)"
  5123   shows "SNa (\<theta>_n,\<theta>_c<M>)"
  5116 using a b c
  5124   using a b c
  5117 proof(nominal_induct avoiding: \<theta>_n \<theta>_c rule: typing.strong_induct)
  5125 proof(nominal_induct avoiding: \<theta>_n \<theta>_c rule: typing.strong_induct)
  5118   case (TAx \<Gamma> \<Delta> x B a \<theta>_n \<theta>_c)
  5126   case (TAx \<Gamma> \<Delta> x B a \<theta>_n \<theta>_c)
  5119   then show ?case
  5127   then show ?case
  5120     apply -
  5128     apply -
  5121     apply(drule ncloses_elim)
  5129     apply(drule ncloses_elim)
  5122     apply(assumption)
  5130      apply(assumption)
  5123     apply(drule ccloses_elim)
  5131     apply(drule ccloses_elim)
  5124     apply(assumption)
  5132      apply(assumption)
  5125     apply(erule exE)+
  5133     apply(erule exE)+
  5126     apply(erule conjE)+
  5134     apply(erule conjE)+
  5127     apply(rule_tac s="Cut <c>.P (xa).Pa" and t="\<theta>_n,\<theta>_c<Ax x a>" in subst)
  5135     apply(rule_tac s="Cut <c>.P (xa).Pa" and t="\<theta>_n,\<theta>_c<Ax x a>" in subst)
  5128     apply(rule sym)
  5136      apply(rule sym)
  5129     apply(simp only: psubst_Ax)
  5137      apply(simp only: psubst_Ax)
  5130     apply(simp add: CUT_SNa)
  5138     apply(simp add: CUT_SNa)
  5131     done
  5139     done
  5132 next
  5140 next
  5133   case (TNotR x \<Gamma> B M \<Delta> \<Delta>' a \<theta>_n \<theta>_c) 
  5141   case (TNotR x \<Gamma> B M \<Delta> \<Delta>' a \<theta>_n \<theta>_c) 
  5134   then show ?case 
  5142   then show ?case 
  5135     apply(simp)
  5143     apply(simp)
  5136     apply(subgoal_tac "(a,NOT B) \<in> set \<Delta>'")
  5144     apply(subgoal_tac "(a,NOT B) \<in> set \<Delta>'")
  5137     apply(drule ccloses_elim)
  5145      apply(drule ccloses_elim)
  5138     apply(assumption)
  5146       apply(assumption)
  5139     apply(erule exE)+
  5147      apply(erule exE)+
  5140     apply(simp)
  5148      apply(simp)
  5141     apply(generate_fresh "coname")
  5149      apply(generate_fresh "coname")
  5142     apply(fresh_fun_simp)
  5150      apply(fresh_fun_simp)
  5143     apply(rule_tac B="NOT B" in CUT_SNa)
  5151      apply(rule_tac B="NOT B" in CUT_SNa)
  5144     apply(simp)
  5152       apply(simp)
  5145     apply(rule disjI2)
  5153       apply(rule disjI2)
  5146     apply(rule disjI2)
  5154       apply(rule disjI2)
  5147     apply(rule_tac x="c" in exI)
  5155       apply(rule_tac x="c" in exI)
  5148     apply(rule_tac x="x" in exI)
  5156       apply(rule_tac x="x" in exI)
  5149     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5157       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5150     apply(simp)
  5158       apply(simp)
  5151     apply(rule conjI)
  5159       apply(rule conjI)
  5152     apply(rule fic.intros)
  5160        apply(rule fic.intros)
  5153     apply(rule psubst_fresh_coname)
  5161        apply(rule psubst_fresh_coname)
  5154     apply(simp)
  5162          apply(simp)
  5155     apply(simp)
  5163         apply(simp)
  5156     apply(simp)
  5164        apply(simp)
  5157     apply(rule BINDING_implies_CAND)
  5165       apply(rule BINDING_implies_CAND)
  5158     apply(unfold BINDINGn_def)
  5166       apply(unfold BINDINGn_def)
  5159     apply(simp)
  5167       apply(simp)
  5160     apply(rule_tac x="x" in exI)
  5168       apply(rule_tac x="x" in exI)
  5161     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5169       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5162     apply(simp)
  5170       apply(simp)
  5163     apply(rule allI)+
  5171       apply(rule allI)+
  5164     apply(rule impI)
  5172       apply(rule impI)
  5165     apply(simp add: psubst_nsubst[symmetric])
  5173       apply(simp add: psubst_nsubst[symmetric])
  5166     apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
  5174       apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
  5167     apply(drule_tac x="\<theta>_c" in meta_spec)
  5175       apply(drule_tac x="\<theta>_c" in meta_spec)
  5168     apply(drule meta_mp)
  5176       apply(drule meta_mp)
  5169     apply(rule ncloses_extend)
  5177        apply(rule ncloses_extend)
  5170     apply(assumption)
  5178           apply(assumption)
  5171     apply(assumption)
  5179          apply(assumption)
  5172     apply(assumption)
  5180         apply(assumption)
  5173     apply(assumption)
  5181        apply(assumption)
  5174     apply(drule meta_mp)
  5182       apply(drule meta_mp)
  5175     apply(rule ccloses_subset)
  5183        apply(rule ccloses_subset)
  5176     apply(assumption)
  5184         apply(assumption)
  5177     apply(blast)
  5185        apply(blast)
  5178     apply(assumption)
  5186       apply(assumption)
  5179     apply(simp)
  5187      apply(simp)
  5180     apply(blast)
  5188     apply(blast)
  5181     done
  5189     done
  5182 next
  5190 next
  5183   case (TNotL a \<Delta> \<Gamma> M B \<Gamma>' x \<theta>_n \<theta>_c)
  5191   case (TNotL a \<Delta> \<Gamma> M B \<Gamma>' x \<theta>_n \<theta>_c)
  5184   then show ?case
  5192   then show ?case
  5185     apply(simp)
  5193     apply(simp)
  5186     apply(subgoal_tac "(x,NOT B) \<in> set \<Gamma>'")
  5194     apply(subgoal_tac "(x,NOT B) \<in> set \<Gamma>'")
  5187     apply(drule ncloses_elim)
  5195      apply(drule ncloses_elim)
  5188     apply(assumption)
  5196       apply(assumption)
  5189     apply(erule exE)+
  5197      apply(erule exE)+
  5190     apply(simp del: NEGc.simps)
  5198      apply(simp del: NEGc.simps)
  5191     apply(generate_fresh "name")
  5199      apply(generate_fresh "name")
  5192     apply(fresh_fun_simp)
  5200      apply(fresh_fun_simp)
  5193     apply(rule_tac B="NOT B" in CUT_SNa)
  5201      apply(rule_tac B="NOT B" in CUT_SNa)
  5194     apply(simp)
  5202       apply(simp)
  5195     apply(rule NEG_intro)
  5203      apply(rule NEG_intro)
  5196     apply(simp (no_asm))
  5204      apply(simp (no_asm))
  5197     apply(rule disjI2)
  5205      apply(rule disjI2)
  5198     apply(rule disjI2)
  5206      apply(rule disjI2)
  5199     apply(rule_tac x="a" in exI)
  5207      apply(rule_tac x="a" in exI)
  5200     apply(rule_tac x="ca" in exI)
  5208      apply(rule_tac x="ca" in exI)
  5201     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5209      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5202     apply(simp del: NEGc.simps)
  5210      apply(simp del: NEGc.simps)
  5203     apply(rule conjI)
  5211      apply(rule conjI)
  5204     apply(rule fin.intros)
  5212       apply(rule fin.intros)
  5205     apply(rule psubst_fresh_name)
  5213       apply(rule psubst_fresh_name)
  5206     apply(simp)
  5214         apply(simp)
  5207     apply(simp)
  5215        apply(simp)
  5208     apply(simp)
  5216       apply(simp)
  5209     apply(rule BINDING_implies_CAND)
  5217      apply(rule BINDING_implies_CAND)
  5210     apply(unfold BINDINGc_def)
  5218      apply(unfold BINDINGc_def)
  5211     apply(simp (no_asm))
  5219      apply(simp (no_asm))
  5212     apply(rule_tac x="a" in exI)
  5220      apply(rule_tac x="a" in exI)
  5213     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5221      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5214     apply(simp (no_asm))
  5222      apply(simp (no_asm))
  5215     apply(rule allI)+
  5223      apply(rule allI)+
  5216     apply(rule impI)
  5224      apply(rule impI)
  5217     apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5225      apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5218     apply(drule_tac x="\<theta>_n" in meta_spec)
  5226      apply(drule_tac x="\<theta>_n" in meta_spec)
  5219     apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5227      apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5220     apply(drule meta_mp)
  5228      apply(drule meta_mp)
  5221     apply(rule ncloses_subset)
  5229       apply(rule ncloses_subset)
  5222     apply(assumption)
  5230        apply(assumption)
  5223     apply(blast)
  5231       apply(blast)
  5224     apply(drule meta_mp)
  5232      apply(drule meta_mp)
  5225     apply(rule ccloses_extend)
  5233       apply(rule ccloses_extend)
  5226     apply(assumption)
  5234          apply(assumption)
  5227     apply(assumption)
  5235         apply(assumption)
  5228     apply(assumption)
  5236        apply(assumption)
  5229     apply(assumption)
  5237       apply(assumption)
  5230     apply(assumption)
  5238      apply(assumption)
  5231     apply(blast)
  5239     apply(blast)
  5232     done
  5240     done
  5233 next
  5241 next
  5234   case (TAndL1 x \<Gamma> y B1 M \<Delta> \<Gamma>' B2 \<theta>_n \<theta>_c)
  5242   case (TAndL1 x \<Gamma> y B1 M \<Delta> \<Gamma>' B2 \<theta>_n \<theta>_c)
  5235   then show ?case     
  5243   then show ?case     
  5236     apply(simp)
  5244     apply(simp)
  5237     apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
  5245     apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
  5238     apply(drule ncloses_elim)
  5246      apply(drule ncloses_elim)
  5239     apply(assumption)
  5247       apply(assumption)
  5240     apply(erule exE)+ 
  5248      apply(erule exE)+ 
  5241     apply(simp del: NEGc.simps)
  5249      apply(simp del: NEGc.simps)
  5242     apply(generate_fresh "name")
  5250      apply(generate_fresh "name")
  5243     apply(fresh_fun_simp)
  5251      apply(fresh_fun_simp)
  5244     apply(rule_tac B="B1 AND B2" in CUT_SNa)
  5252      apply(rule_tac B="B1 AND B2" in CUT_SNa)
  5245     apply(simp)
  5253       apply(simp)
  5246     apply(rule NEG_intro)
  5254      apply(rule NEG_intro)
  5247     apply(simp (no_asm))
  5255      apply(simp (no_asm))
  5248     apply(rule disjI2)
  5256      apply(rule disjI2)
  5249     apply(rule disjI2)
  5257      apply(rule disjI2)
  5250     apply(rule disjI1)
  5258      apply(rule disjI1)
  5251     apply(rule_tac x="x" in exI)
  5259      apply(rule_tac x="x" in exI)
  5252     apply(rule_tac x="ca" in exI)
  5260      apply(rule_tac x="ca" in exI)
  5253     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5261      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5254     apply(simp del: NEGc.simps)
  5262      apply(simp del: NEGc.simps)
  5255     apply(rule conjI)
  5263      apply(rule conjI)
  5256     apply(rule fin.intros)
  5264       apply(rule fin.intros)
  5257     apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5265       apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5258     apply(rule psubst_fresh_name)
  5266       apply(rule psubst_fresh_name)
  5259     apply(simp)
  5267         apply(simp)
  5260     apply(simp)
  5268        apply(simp)
  5261     apply(simp)
  5269       apply(simp)
  5262     apply(rule BINDING_implies_CAND)
  5270      apply(rule BINDING_implies_CAND)
  5263     apply(unfold BINDINGn_def)
  5271      apply(unfold BINDINGn_def)
  5264     apply(simp (no_asm))
  5272      apply(simp (no_asm))
  5265     apply(rule_tac x="x" in exI)
  5273      apply(rule_tac x="x" in exI)
  5266     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5274      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5267     apply(simp (no_asm))
  5275      apply(simp (no_asm))
  5268     apply(rule allI)+
  5276      apply(rule allI)+
  5269     apply(rule impI)
  5277      apply(rule impI)
  5270     apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5278      apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5271     apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
  5279      apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
  5272     apply(drule_tac x="\<theta>_c" in meta_spec)
  5280      apply(drule_tac x="\<theta>_c" in meta_spec)
  5273     apply(drule meta_mp)
  5281      apply(drule meta_mp)
  5274     apply(rule ncloses_extend)
  5282       apply(rule ncloses_extend)
  5275     apply(rule ncloses_subset)
  5283          apply(rule ncloses_subset)
  5276     apply(assumption)
  5284           apply(assumption)
  5277     apply(blast)
  5285          apply(blast)
  5278     apply(simp)
  5286         apply(simp)
  5279     apply(simp)
  5287        apply(simp)
  5280     apply(simp)
  5288       apply(simp)
  5281     apply(drule meta_mp)
  5289      apply(drule meta_mp)
  5282     apply(assumption)
  5290       apply(assumption)
  5283     apply(assumption)
  5291      apply(assumption)
  5284     apply(blast)
  5292     apply(blast)
  5285     done
  5293     done
  5286 next
  5294 next
  5287   case (TAndL2 x \<Gamma> y B2 M \<Delta> \<Gamma>' B1 \<theta>_n \<theta>_c)
  5295   case (TAndL2 x \<Gamma> y B2 M \<Delta> \<Gamma>' B1 \<theta>_n \<theta>_c)
  5288   then show ?case 
  5296   then show ?case 
  5289     apply(simp)
  5297     apply(simp)
  5290     apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
  5298     apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
  5291     apply(drule ncloses_elim)
  5299      apply(drule ncloses_elim)
  5292     apply(assumption)
  5300       apply(assumption)
  5293     apply(erule exE)+ 
  5301      apply(erule exE)+ 
  5294     apply(simp del: NEGc.simps)
  5302      apply(simp del: NEGc.simps)
  5295     apply(generate_fresh "name")
  5303      apply(generate_fresh "name")
  5296     apply(fresh_fun_simp)
  5304      apply(fresh_fun_simp)
  5297     apply(rule_tac B="B1 AND B2" in CUT_SNa)
  5305      apply(rule_tac B="B1 AND B2" in CUT_SNa)
  5298     apply(simp)
  5306       apply(simp)
  5299     apply(rule NEG_intro)
  5307      apply(rule NEG_intro)
  5300     apply(simp (no_asm))
  5308      apply(simp (no_asm))
  5301     apply(rule disjI2)
  5309      apply(rule disjI2)
  5302     apply(rule disjI2)
  5310      apply(rule disjI2)
  5303     apply(rule disjI2)
  5311      apply(rule disjI2)
  5304     apply(rule_tac x="x" in exI)
  5312      apply(rule_tac x="x" in exI)
  5305     apply(rule_tac x="ca" in exI)
  5313      apply(rule_tac x="ca" in exI)
  5306     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5314      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5307     apply(simp del: NEGc.simps)
  5315      apply(simp del: NEGc.simps)
  5308     apply(rule conjI)
  5316      apply(rule conjI)
  5309     apply(rule fin.intros)
  5317       apply(rule fin.intros)
  5310     apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5318       apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5311     apply(rule psubst_fresh_name)
  5319       apply(rule psubst_fresh_name)
  5312     apply(simp)
  5320         apply(simp)
  5313     apply(simp)
  5321        apply(simp)
  5314     apply(simp)
  5322       apply(simp)
  5315     apply(rule BINDING_implies_CAND)
  5323      apply(rule BINDING_implies_CAND)
  5316     apply(unfold BINDINGn_def)
  5324      apply(unfold BINDINGn_def)
  5317     apply(simp (no_asm))
  5325      apply(simp (no_asm))
  5318     apply(rule_tac x="x" in exI)
  5326      apply(rule_tac x="x" in exI)
  5319     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5327      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5320     apply(simp (no_asm))
  5328      apply(simp (no_asm))
  5321     apply(rule allI)+
  5329      apply(rule allI)+
  5322     apply(rule impI)
  5330      apply(rule impI)
  5323     apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5331      apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5324     apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
  5332      apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
  5325     apply(drule_tac x="\<theta>_c" in meta_spec)
  5333      apply(drule_tac x="\<theta>_c" in meta_spec)
  5326     apply(drule meta_mp)
  5334      apply(drule meta_mp)
  5327     apply(rule ncloses_extend)
  5335       apply(rule ncloses_extend)
  5328     apply(rule ncloses_subset)
  5336          apply(rule ncloses_subset)
  5329     apply(assumption)
  5337           apply(assumption)
  5330     apply(blast)
  5338          apply(blast)
  5331     apply(simp)
  5339         apply(simp)
  5332     apply(simp)
  5340        apply(simp)
  5333     apply(simp)
  5341       apply(simp)
  5334     apply(drule meta_mp)
  5342      apply(drule meta_mp)
  5335     apply(assumption)
  5343       apply(assumption)
  5336     apply(assumption)
  5344      apply(assumption)
  5337     apply(blast)
  5345     apply(blast)
  5338     done
  5346     done
  5339 next
  5347 next
  5340   case (TAndR a \<Delta> N c b M \<Gamma> B C \<Delta>' \<theta>_n \<theta>_c)
  5348   case (TAndR a \<Delta> N c b M \<Gamma> B C \<Delta>' \<theta>_n \<theta>_c)
  5341   then show ?case 
  5349   then show ?case 
  5342     apply(simp)
  5350     apply(simp)
  5343     apply(subgoal_tac "(c,B AND C) \<in> set \<Delta>'")
  5351     apply(subgoal_tac "(c,B AND C) \<in> set \<Delta>'")
  5344     apply(drule ccloses_elim)
  5352      apply(drule ccloses_elim)
  5345     apply(assumption)
  5353       apply(assumption)
  5346     apply(erule exE)+
  5354      apply(erule exE)+
  5347     apply(simp)
  5355      apply(simp)
  5348     apply(generate_fresh "coname")
  5356      apply(generate_fresh "coname")
  5349     apply(fresh_fun_simp)
  5357      apply(fresh_fun_simp)
  5350     apply(rule_tac B="B AND C" in CUT_SNa)
  5358      apply(rule_tac B="B AND C" in CUT_SNa)
  5351     apply(simp)
  5359       apply(simp)
  5352     apply(rule disjI2)
  5360       apply(rule disjI2)
  5353     apply(rule disjI2)
  5361       apply(rule disjI2)
  5354     apply(rule_tac x="ca" in exI)
  5362       apply(rule_tac x="ca" in exI)
  5355     apply(rule_tac x="a" in exI)
  5363       apply(rule_tac x="a" in exI)
  5356     apply(rule_tac x="b" in exI)
  5364       apply(rule_tac x="b" in exI)
  5357     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5365       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5358     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5366       apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5359     apply(simp)
  5367       apply(simp)
  5360     apply(rule conjI)
  5368       apply(rule conjI)
  5361     apply(rule fic.intros)
  5369        apply(rule fic.intros)
  5362     apply(simp add: abs_fresh fresh_prod fresh_atm)
  5370         apply(simp add: abs_fresh fresh_prod fresh_atm)
  5363     apply(rule psubst_fresh_coname)
  5371         apply(rule psubst_fresh_coname)
  5364     apply(simp)
  5372           apply(simp)
  5365     apply(simp)
  5373          apply(simp)
  5366     apply(simp)
  5374         apply(simp)
  5367     apply(simp add: abs_fresh fresh_prod fresh_atm)
  5375        apply(simp add: abs_fresh fresh_prod fresh_atm)
  5368     apply(rule psubst_fresh_coname)
  5376        apply(rule psubst_fresh_coname)
  5369     apply(simp)
  5377          apply(simp)
  5370     apply(simp)
  5378         apply(simp)
  5371     apply(simp)
  5379        apply(simp)
  5372     apply(rule conjI)
  5380       apply(rule conjI)
  5373     apply(rule BINDING_implies_CAND)
  5381        apply(rule BINDING_implies_CAND)
  5374     apply(unfold BINDINGc_def)
  5382        apply(unfold BINDINGc_def)
  5375     apply(simp)
  5383        apply(simp)
  5376     apply(rule_tac x="a" in exI)
  5384        apply(rule_tac x="a" in exI)
  5377     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5385        apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5378     apply(simp)
  5386        apply(simp)
  5379     apply(rule allI)+
  5387        apply(rule allI)+
  5380     apply(rule impI)
  5388        apply(rule impI)
  5381     apply(simp add: psubst_csubst[symmetric])
  5389        apply(simp add: psubst_csubst[symmetric])
  5382     apply(drule_tac x="\<theta>_n" in meta_spec)
  5390        apply(drule_tac x="\<theta>_n" in meta_spec)
  5383     apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5391        apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5384     apply(drule meta_mp)
  5392        apply(drule meta_mp)
  5385     apply(assumption)
  5393         apply(assumption)
  5386     apply(drule meta_mp)
  5394        apply(drule meta_mp)
  5387     apply(rule ccloses_extend)
  5395         apply(rule ccloses_extend)
  5388     apply(rule ccloses_subset)
  5396            apply(rule ccloses_subset)
  5389     apply(assumption)
  5397             apply(assumption)
  5390     apply(blast)
  5398            apply(blast)
  5391     apply(simp)
  5399           apply(simp)
  5392     apply(simp)
  5400          apply(simp)
  5393     apply(assumption)
  5401         apply(assumption)
  5394     apply(assumption)
  5402        apply(assumption)
  5395     apply(rule BINDING_implies_CAND)
  5403       apply(rule BINDING_implies_CAND)
  5396     apply(unfold BINDINGc_def)
  5404       apply(unfold BINDINGc_def)
  5397     apply(simp)
  5405       apply(simp)
  5398     apply(rule_tac x="b" in exI)
  5406       apply(rule_tac x="b" in exI)
  5399     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5407       apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5400     apply(simp)
  5408       apply(simp)
  5401     apply(rule allI)+
  5409       apply(rule allI)+
  5402     apply(rule impI)
  5410       apply(rule impI)
  5403     apply(simp add: psubst_csubst[symmetric])
  5411       apply(simp add: psubst_csubst[symmetric])
  5404     apply(rotate_tac 14)
  5412       apply(rotate_tac 14)
  5405     apply(drule_tac x="\<theta>_n" in meta_spec)
  5413       apply(drule_tac x="\<theta>_n" in meta_spec)
  5406     apply(drule_tac x="(b,xa,Pa)#\<theta>_c" in meta_spec)
  5414       apply(drule_tac x="(b,xa,Pa)#\<theta>_c" in meta_spec)
  5407     apply(drule meta_mp)
  5415       apply(drule meta_mp)
  5408     apply(assumption)
  5416        apply(assumption)
  5409     apply(drule meta_mp)
  5417       apply(drule meta_mp)
  5410     apply(rule ccloses_extend)
  5418        apply(rule ccloses_extend)
  5411     apply(rule ccloses_subset)
  5419           apply(rule ccloses_subset)
  5412     apply(assumption)
  5420            apply(assumption)
  5413     apply(blast)
  5421           apply(blast)
  5414     apply(simp)
  5422          apply(simp)
  5415     apply(simp)
  5423         apply(simp)
  5416     apply(assumption)
  5424        apply(assumption)
  5417     apply(assumption)
  5425       apply(assumption)
  5418     apply(simp)
  5426      apply(simp)
  5419     apply(blast)
  5427     apply(blast)
  5420     done
  5428     done
  5421 next
  5429 next
  5422   case (TOrL x \<Gamma> N z y M B \<Delta> C \<Gamma>' \<theta>_n \<theta>_c)
  5430   case (TOrL x \<Gamma> N z y M B \<Delta> C \<Gamma>' \<theta>_n \<theta>_c)
  5423   then show ?case 
  5431   then show ?case 
  5424     apply(simp)
  5432     apply(simp)
  5425     apply(subgoal_tac "(z,B OR C) \<in> set \<Gamma>'")
  5433     apply(subgoal_tac "(z,B OR C) \<in> set \<Gamma>'")
  5426     apply(drule ncloses_elim)
  5434      apply(drule ncloses_elim)
  5427     apply(assumption)
  5435       apply(assumption)
  5428     apply(erule exE)+
  5436      apply(erule exE)+
  5429     apply(simp del: NEGc.simps)
  5437      apply(simp del: NEGc.simps)
  5430     apply(generate_fresh "name")
  5438      apply(generate_fresh "name")
  5431     apply(fresh_fun_simp)
  5439      apply(fresh_fun_simp)
  5432     apply(rule_tac B="B OR C" in CUT_SNa)
  5440      apply(rule_tac B="B OR C" in CUT_SNa)
  5433     apply(simp)
  5441       apply(simp)
  5434     apply(rule NEG_intro)
  5442      apply(rule NEG_intro)
  5435     apply(simp (no_asm))
  5443      apply(simp (no_asm))
  5436     apply(rule disjI2)
  5444      apply(rule disjI2)
  5437     apply(rule disjI2)
  5445      apply(rule disjI2)
  5438     apply(rule_tac x="x" in exI)
  5446      apply(rule_tac x="x" in exI)
  5439     apply(rule_tac x="y" in exI)
  5447      apply(rule_tac x="y" in exI)
  5440     apply(rule_tac x="ca" in exI)
  5448      apply(rule_tac x="ca" in exI)
  5441     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5449      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5442     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5450      apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5443     apply(simp del: NEGc.simps)
  5451      apply(simp del: NEGc.simps)
  5444     apply(rule conjI)
  5452      apply(rule conjI)
  5445     apply(rule fin.intros)
  5453       apply(rule fin.intros)
  5446     apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5454        apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5447     apply(rule psubst_fresh_name)
  5455        apply(rule psubst_fresh_name)
  5448     apply(simp)
  5456          apply(simp)
  5449     apply(simp)
  5457         apply(simp)
  5450     apply(simp)
  5458        apply(simp)
  5451     apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5459       apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5452     apply(rule psubst_fresh_name)
  5460       apply(rule psubst_fresh_name)
  5453     apply(simp)
  5461         apply(simp)
  5454     apply(simp)
  5462        apply(simp)
  5455     apply(simp)
  5463       apply(simp)
  5456     apply(rule conjI)
  5464      apply(rule conjI)
  5457     apply(rule BINDING_implies_CAND)
  5465       apply(rule BINDING_implies_CAND)
  5458     apply(unfold BINDINGn_def)
  5466       apply(unfold BINDINGn_def)
  5459     apply(simp del: NEGc.simps)
  5467       apply(simp del: NEGc.simps)
  5460     apply(rule_tac x="x" in exI)
  5468       apply(rule_tac x="x" in exI)
  5461     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5469       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5462     apply(simp del: NEGc.simps)
  5470       apply(simp del: NEGc.simps)
  5463     apply(rule allI)+
  5471       apply(rule allI)+
  5464     apply(rule impI)
  5472       apply(rule impI)
  5465     apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5473       apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5466     apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
  5474       apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
  5467     apply(drule_tac x="\<theta>_c" in meta_spec)
  5475       apply(drule_tac x="\<theta>_c" in meta_spec)
  5468     apply(drule meta_mp)
  5476       apply(drule meta_mp)
  5469     apply(rule ncloses_extend)
  5477        apply(rule ncloses_extend)
  5470     apply(rule ncloses_subset)
  5478           apply(rule ncloses_subset)
  5471     apply(assumption)
  5479            apply(assumption)
  5472     apply(blast)
  5480           apply(blast)
  5473     apply(simp)
  5481          apply(simp)
  5474     apply(simp)
  5482         apply(simp)
  5475     apply(assumption)
  5483        apply(assumption)
  5476     apply(drule meta_mp)
  5484       apply(drule meta_mp)
  5477     apply(assumption)
  5485        apply(assumption)
  5478     apply(assumption)
  5486       apply(assumption)
  5479     apply(rule BINDING_implies_CAND)
  5487      apply(rule BINDING_implies_CAND)
  5480     apply(unfold BINDINGn_def)
  5488      apply(unfold BINDINGn_def)
  5481     apply(simp del: NEGc.simps)
  5489      apply(simp del: NEGc.simps)
  5482     apply(rule_tac x="y" in exI)
  5490      apply(rule_tac x="y" in exI)
  5483     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5491      apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5484     apply(simp del: NEGc.simps)
  5492      apply(simp del: NEGc.simps)
  5485     apply(rule allI)+
  5493      apply(rule allI)+
  5486     apply(rule impI)
  5494      apply(rule impI)
  5487     apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5495      apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5488     apply(rotate_tac 14)
  5496      apply(rotate_tac 14)
  5489     apply(drule_tac x="(y,a,Pa)#\<theta>_n" in meta_spec)
  5497      apply(drule_tac x="(y,a,Pa)#\<theta>_n" in meta_spec)
  5490     apply(drule_tac x="\<theta>_c" in meta_spec)
  5498      apply(drule_tac x="\<theta>_c" in meta_spec)
  5491     apply(drule meta_mp)
  5499      apply(drule meta_mp)
  5492     apply(rule ncloses_extend)
  5500       apply(rule ncloses_extend)
  5493     apply(rule ncloses_subset)
  5501          apply(rule ncloses_subset)
  5494     apply(assumption)
  5502           apply(assumption)
  5495     apply(blast)
  5503          apply(blast)
  5496     apply(simp)
  5504         apply(simp)
  5497     apply(simp)
  5505        apply(simp)
  5498     apply(assumption)
  5506       apply(assumption)
  5499     apply(drule meta_mp)
  5507      apply(drule meta_mp)
  5500     apply(assumption)
  5508       apply(assumption)
  5501     apply(assumption)
  5509      apply(assumption)
  5502     apply(blast)
  5510     apply(blast)
  5503     done
  5511     done
  5504 next
  5512 next
  5505   case (TOrR1 a \<Delta> b \<Gamma> M B1 \<Delta>' B2 \<theta>_n \<theta>_c)
  5513   case (TOrR1 a \<Delta> b \<Gamma> M B1 \<Delta>' B2 \<theta>_n \<theta>_c)
  5506   then show ?case
  5514   then show ?case
  5507     apply(simp)
  5515     apply(simp)
  5508     apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
  5516     apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
  5509     apply(drule ccloses_elim)
  5517      apply(drule ccloses_elim)
  5510     apply(assumption)
  5518       apply(assumption)
  5511     apply(erule exE)+ 
  5519      apply(erule exE)+ 
  5512     apply(simp del: NEGc.simps)
  5520      apply(simp del: NEGc.simps)
  5513     apply(generate_fresh "coname")
  5521      apply(generate_fresh "coname")
  5514     apply(fresh_fun_simp)
  5522      apply(fresh_fun_simp)
  5515     apply(rule_tac B="B1 OR B2" in CUT_SNa)
  5523      apply(rule_tac B="B1 OR B2" in CUT_SNa)
  5516     apply(simp)
  5524       apply(simp)
  5517     apply(rule disjI2)
  5525       apply(rule disjI2)
  5518     apply(rule disjI2)
  5526       apply(rule disjI2)
  5519     apply(rule disjI1)
  5527       apply(rule disjI1)
  5520     apply(rule_tac x="a" in exI)
  5528       apply(rule_tac x="a" in exI)
  5521     apply(rule_tac x="c" in exI)
  5529       apply(rule_tac x="c" in exI)
  5522     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5530       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5523     apply(simp)
  5531       apply(simp)
  5524     apply(rule conjI)
  5532       apply(rule conjI)
  5525     apply(rule fic.intros)
  5533        apply(rule fic.intros)
  5526     apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5534        apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5527     apply(rule psubst_fresh_coname)
  5535        apply(rule psubst_fresh_coname)
  5528     apply(simp)
  5536          apply(simp)
  5529     apply(simp)
  5537         apply(simp)
  5530     apply(simp)
  5538        apply(simp)
  5531     apply(rule BINDING_implies_CAND)
  5539       apply(rule BINDING_implies_CAND)
  5532     apply(unfold BINDINGc_def)
  5540       apply(unfold BINDINGc_def)
  5533     apply(simp (no_asm))
  5541       apply(simp (no_asm))
  5534     apply(rule_tac x="a" in exI)
  5542       apply(rule_tac x="a" in exI)
  5535     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5543       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5536     apply(simp (no_asm))
  5544       apply(simp (no_asm))
  5537     apply(rule allI)+
  5545       apply(rule allI)+
  5538     apply(rule impI)    
  5546       apply(rule impI)    
  5539     apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5547       apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5540     apply(drule_tac x="\<theta>_n" in meta_spec)
  5548       apply(drule_tac x="\<theta>_n" in meta_spec)
  5541     apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5549       apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5542     apply(drule meta_mp)
  5550       apply(drule meta_mp)
  5543     apply(assumption)
  5551        apply(assumption)
  5544     apply(drule meta_mp)
  5552       apply(drule meta_mp)
  5545     apply(rule ccloses_extend)
  5553        apply(rule ccloses_extend)
  5546     apply(rule ccloses_subset)
  5554           apply(rule ccloses_subset)
  5547     apply(assumption)
  5555            apply(assumption)
  5548     apply(blast)
  5556           apply(blast)
  5549     apply(simp)
  5557          apply(simp)
  5550     apply(simp)
  5558         apply(simp)
  5551     apply(simp)
  5559        apply(simp)
  5552     apply(assumption)
  5560       apply(assumption)
  5553     apply(simp)
  5561      apply(simp)
  5554     apply(blast)
  5562     apply(blast)
  5555     done
  5563     done
  5556 next
  5564 next
  5557   case (TOrR2 a \<Delta> b \<Gamma> M B2 \<Delta>' B1 \<theta>_n \<theta>_c)
  5565   case (TOrR2 a \<Delta> b \<Gamma> M B2 \<Delta>' B1 \<theta>_n \<theta>_c)
  5558   then show ?case 
  5566   then show ?case 
  5559     apply(simp)
  5567     apply(simp)
  5560     apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
  5568     apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
  5561     apply(drule ccloses_elim)
  5569      apply(drule ccloses_elim)
  5562     apply(assumption)
  5570       apply(assumption)
  5563     apply(erule exE)+ 
  5571      apply(erule exE)+ 
  5564     apply(simp del: NEGc.simps)
  5572      apply(simp del: NEGc.simps)
  5565     apply(generate_fresh "coname")
  5573      apply(generate_fresh "coname")
  5566     apply(fresh_fun_simp)
  5574      apply(fresh_fun_simp)
  5567     apply(rule_tac B="B1 OR B2" in CUT_SNa)
  5575      apply(rule_tac B="B1 OR B2" in CUT_SNa)
  5568     apply(simp)
  5576       apply(simp)
  5569     apply(rule disjI2)
  5577       apply(rule disjI2)
  5570     apply(rule disjI2)
  5578       apply(rule disjI2)
  5571     apply(rule disjI2)
  5579       apply(rule disjI2)
  5572     apply(rule_tac x="a" in exI)
  5580       apply(rule_tac x="a" in exI)
  5573     apply(rule_tac x="c" in exI)
  5581       apply(rule_tac x="c" in exI)
  5574     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5582       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5575     apply(simp)
  5583       apply(simp)
  5576     apply(rule conjI)
  5584       apply(rule conjI)
  5577     apply(rule fic.intros)
  5585        apply(rule fic.intros)
  5578     apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5586        apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5579     apply(rule psubst_fresh_coname)
  5587        apply(rule psubst_fresh_coname)
  5580     apply(simp)
  5588          apply(simp)
  5581     apply(simp)
  5589         apply(simp)
  5582     apply(simp)
  5590        apply(simp)
  5583     apply(rule BINDING_implies_CAND)
  5591       apply(rule BINDING_implies_CAND)
  5584     apply(unfold BINDINGc_def)
  5592       apply(unfold BINDINGc_def)
  5585     apply(simp (no_asm))
  5593       apply(simp (no_asm))
  5586     apply(rule_tac x="a" in exI)
  5594       apply(rule_tac x="a" in exI)
  5587     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5595       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5588     apply(simp (no_asm))
  5596       apply(simp (no_asm))
  5589     apply(rule allI)+
  5597       apply(rule allI)+
  5590     apply(rule impI)    
  5598       apply(rule impI)    
  5591     apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5599       apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5592     apply(drule_tac x="\<theta>_n" in meta_spec)
  5600       apply(drule_tac x="\<theta>_n" in meta_spec)
  5593     apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5601       apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5594     apply(drule meta_mp)
  5602       apply(drule meta_mp)
  5595     apply(assumption)
  5603        apply(assumption)
  5596     apply(drule meta_mp)
  5604       apply(drule meta_mp)
  5597     apply(rule ccloses_extend)
  5605        apply(rule ccloses_extend)
  5598     apply(rule ccloses_subset)
  5606           apply(rule ccloses_subset)
  5599     apply(assumption)
  5607            apply(assumption)
  5600     apply(blast)
  5608           apply(blast)
  5601     apply(simp)
  5609          apply(simp)
  5602     apply(simp)
  5610         apply(simp)
  5603     apply(simp)
  5611        apply(simp)
  5604     apply(assumption)
  5612       apply(assumption)
  5605     apply(simp)
  5613      apply(simp)
  5606     apply(blast)
  5614     apply(blast)
  5607     done
  5615     done
  5608 next
  5616 next
  5609   case (TImpL a \<Delta> N x \<Gamma> M y B C \<Gamma>' \<theta>_n \<theta>_c)
  5617   case (TImpL a \<Delta> N x \<Gamma> M y B C \<Gamma>' \<theta>_n \<theta>_c)
  5610   then show ?case
  5618   then show ?case
  5611     apply(simp)
  5619     apply(simp)
  5612     apply(subgoal_tac "(y,B IMP C) \<in> set \<Gamma>'")
  5620     apply(subgoal_tac "(y,B IMP C) \<in> set \<Gamma>'")
  5613     apply(drule ncloses_elim)
  5621      apply(drule ncloses_elim)
  5614     apply(assumption)
  5622       apply(assumption)
  5615     apply(erule exE)+
  5623      apply(erule exE)+
  5616     apply(simp del: NEGc.simps)
  5624      apply(simp del: NEGc.simps)
  5617     apply(generate_fresh "name")
  5625      apply(generate_fresh "name")
  5618     apply(fresh_fun_simp)
  5626      apply(fresh_fun_simp)
  5619     apply(rule_tac B="B IMP C" in CUT_SNa)
  5627      apply(rule_tac B="B IMP C" in CUT_SNa)
  5620     apply(simp)
  5628       apply(simp)
  5621     apply(rule NEG_intro)
  5629      apply(rule NEG_intro)
  5622     apply(simp (no_asm))
  5630      apply(simp (no_asm))
  5623     apply(rule disjI2)
  5631      apply(rule disjI2)
  5624     apply(rule disjI2)
  5632      apply(rule disjI2)
  5625     apply(rule_tac x="x" in exI)
  5633      apply(rule_tac x="x" in exI)
  5626     apply(rule_tac x="a" in exI)
  5634      apply(rule_tac x="a" in exI)
  5627     apply(rule_tac x="ca" in exI)
  5635      apply(rule_tac x="ca" in exI)
  5628     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5636      apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5629     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5637      apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5630     apply(simp del: NEGc.simps)
  5638      apply(simp del: NEGc.simps)
  5631     apply(rule conjI)
  5639      apply(rule conjI)
  5632     apply(rule fin.intros)
  5640       apply(rule fin.intros)
  5633     apply(rule psubst_fresh_name)
  5641        apply(rule psubst_fresh_name)
  5634     apply(simp)
  5642          apply(simp)
  5635     apply(simp)
  5643         apply(simp)
  5636     apply(simp)
  5644        apply(simp)
  5637     apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5645       apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
  5638     apply(rule psubst_fresh_name)
  5646       apply(rule psubst_fresh_name)
  5639     apply(simp)
  5647         apply(simp)
  5640     apply(simp)
  5648        apply(simp)
  5641     apply(simp)
  5649       apply(simp)
  5642     apply(rule conjI)
  5650      apply(rule conjI)
  5643     apply(rule BINDING_implies_CAND)
  5651       apply(rule BINDING_implies_CAND)
  5644     apply(unfold BINDINGc_def)
  5652       apply(unfold BINDINGc_def)
  5645     apply(simp del: NEGc.simps)
  5653       apply(simp del: NEGc.simps)
  5646     apply(rule_tac x="a" in exI)
  5654       apply(rule_tac x="a" in exI)
  5647     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5655       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5648     apply(simp del: NEGc.simps)
  5656       apply(simp del: NEGc.simps)
  5649     apply(rule allI)+
  5657       apply(rule allI)+
  5650     apply(rule impI)
  5658       apply(rule impI)
  5651     apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5659       apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
  5652     apply(drule_tac x="\<theta>_n" in meta_spec)
  5660       apply(drule_tac x="\<theta>_n" in meta_spec)
  5653     apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5661       apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
  5654     apply(drule meta_mp)
  5662       apply(drule meta_mp)
  5655     apply(rule ncloses_subset)
  5663        apply(rule ncloses_subset)
  5656     apply(assumption)
  5664         apply(assumption)
  5657     apply(blast)
  5665        apply(blast)
  5658     apply(drule meta_mp)
  5666       apply(drule meta_mp)
  5659     apply(rule ccloses_extend)
  5667        apply(rule ccloses_extend)
  5660     apply(assumption)
  5668           apply(assumption)
  5661     apply(simp)
  5669          apply(simp)
  5662     apply(simp)
  5670         apply(simp)
  5663     apply(assumption)
  5671        apply(assumption)
  5664     apply(assumption)
  5672       apply(assumption)
  5665     apply(rule BINDING_implies_CAND)
  5673      apply(rule BINDING_implies_CAND)
  5666     apply(unfold BINDINGn_def)
  5674      apply(unfold BINDINGn_def)
  5667     apply(simp del: NEGc.simps)
  5675      apply(simp del: NEGc.simps)
  5668     apply(rule_tac x="x" in exI)
  5676      apply(rule_tac x="x" in exI)
  5669     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5677      apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5670     apply(simp del: NEGc.simps)
  5678      apply(simp del: NEGc.simps)
  5671     apply(rule allI)+
  5679      apply(rule allI)+
  5672     apply(rule impI)
  5680      apply(rule impI)
  5673     apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5681      apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
  5674     apply(rotate_tac 12)
  5682      apply(rotate_tac 12)
  5675     apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
  5683      apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
  5676     apply(drule_tac x="\<theta>_c" in meta_spec)
  5684      apply(drule_tac x="\<theta>_c" in meta_spec)
  5677     apply(drule meta_mp)
  5685      apply(drule meta_mp)
  5678     apply(rule ncloses_extend)
  5686       apply(rule ncloses_extend)
  5679     apply(rule ncloses_subset)
  5687          apply(rule ncloses_subset)
  5680     apply(assumption)
  5688           apply(assumption)
  5681     apply(blast)
  5689          apply(blast)
  5682     apply(simp)
  5690         apply(simp)
  5683     apply(simp)
  5691        apply(simp)
  5684     apply(assumption)
  5692       apply(assumption)
  5685     apply(drule meta_mp)
  5693      apply(drule meta_mp)
  5686     apply(assumption)
  5694       apply(assumption)
  5687     apply(assumption)
  5695      apply(assumption)
  5688     apply(blast)
  5696     apply(blast)
  5689     done
  5697     done
  5690 next
  5698 next
  5691   case (TImpR a \<Delta> b x \<Gamma> B M C \<Delta>' \<theta>_n \<theta>_c)
  5699   case (TImpR a \<Delta> b x \<Gamma> B M C \<Delta>' \<theta>_n \<theta>_c)
  5692   then show ?case
  5700   then show ?case
  5693     apply(simp)
  5701     apply(simp)
  5694     apply(subgoal_tac "(b,B IMP C) \<in> set \<Delta>'")
  5702     apply(subgoal_tac "(b,B IMP C) \<in> set \<Delta>'")
  5695     apply(drule ccloses_elim)
  5703      apply(drule ccloses_elim)
  5696     apply(assumption)
  5704       apply(assumption)
  5697     apply(erule exE)+
  5705      apply(erule exE)+
  5698     apply(simp)
  5706      apply(simp)
  5699     apply(generate_fresh "coname")
  5707      apply(generate_fresh "coname")
  5700     apply(fresh_fun_simp)
  5708      apply(fresh_fun_simp)
  5701     apply(rule_tac B="B IMP C" in CUT_SNa)
  5709      apply(rule_tac B="B IMP C" in CUT_SNa)
  5702     apply(simp)
  5710       apply(simp)
  5703     apply(rule disjI2)
  5711       apply(rule disjI2)
  5704     apply(rule disjI2)
  5712       apply(rule disjI2)
  5705     apply(rule_tac x="x" in exI)
  5713       apply(rule_tac x="x" in exI)
  5706     apply(rule_tac x="a" in exI)
  5714       apply(rule_tac x="a" in exI)
  5707     apply(rule_tac x="c" in exI)
  5715       apply(rule_tac x="c" in exI)
  5708     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5716       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5709     apply(simp)
  5717       apply(simp)
  5710     apply(rule conjI)
  5718       apply(rule conjI)
  5711     apply(rule fic.intros)
  5719        apply(rule fic.intros)
  5712     apply(simp add: abs_fresh fresh_prod fresh_atm)
  5720        apply(simp add: abs_fresh fresh_prod fresh_atm)
  5713     apply(rule psubst_fresh_coname)
  5721        apply(rule psubst_fresh_coname)
  5714     apply(simp)
  5722          apply(simp)
  5715     apply(simp)
  5723         apply(simp)
  5716     apply(simp)
  5724        apply(simp)
  5717     apply(rule conjI)
  5725       apply(rule conjI)
  5718     apply(rule allI)+
  5726        apply(rule allI)+
  5719     apply(rule impI)
  5727        apply(rule impI)
  5720     apply(simp add: psubst_csubst[symmetric])
  5728        apply(simp add: psubst_csubst[symmetric])
  5721     apply(rule BINDING_implies_CAND)
  5729        apply(rule BINDING_implies_CAND)
  5722     apply(unfold BINDINGn_def)
  5730        apply(unfold BINDINGn_def)
  5723     apply(simp)
  5731        apply(simp)
  5724     apply(rule_tac x="x" in exI)
  5732        apply(rule_tac x="x" in exI)
  5725     apply(rule_tac x="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>" in exI)
  5733        apply(rule_tac x="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>" in exI)
  5726     apply(simp)
  5734        apply(simp)
  5727     apply(rule allI)+
  5735        apply(rule allI)+
  5728     apply(rule impI)
  5736        apply(rule impI)
  5729     apply(rule_tac t="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>{x:=<aa>.Pb}" and 
  5737        apply(rule_tac t="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>{x:=<aa>.Pb}" and 
  5730                    s="((x,aa,Pb)#\<theta>_n),((a,z,Pa)#\<theta>_c)<M>" in subst)
  5738         s="((x,aa,Pb)#\<theta>_n),((a,z,Pa)#\<theta>_c)<M>" in subst)
  5731     apply(rule psubst_nsubst)
  5739         apply(rule psubst_nsubst)
  5732     apply(simp add: fresh_prod fresh_atm fresh_list_cons)
  5740         apply(simp add: fresh_prod fresh_atm fresh_list_cons)
  5733     apply(drule_tac x="(x,aa,Pb)#\<theta>_n" in meta_spec)
  5741        apply(drule_tac x="(x,aa,Pb)#\<theta>_n" in meta_spec)
  5734     apply(drule_tac x="(a,z,Pa)#\<theta>_c" in meta_spec)
  5742        apply(drule_tac x="(a,z,Pa)#\<theta>_c" in meta_spec)
  5735     apply(drule meta_mp)
  5743        apply(drule meta_mp)
  5736     apply(rule ncloses_extend)
  5744         apply(rule ncloses_extend)
  5737     apply(assumption)
  5745            apply(assumption)
  5738     apply(simp)
  5746           apply(simp)
  5739     apply(simp)
  5747          apply(simp)
  5740     apply(simp)
  5748         apply(simp)
  5741     apply(drule meta_mp)
  5749        apply(drule meta_mp)
  5742     apply(rule ccloses_extend)
  5750         apply(rule ccloses_extend)
  5743     apply(rule ccloses_subset)
  5751            apply(rule ccloses_subset)
  5744     apply(assumption)
  5752             apply(assumption)
  5745     apply(blast)
  5753            apply(blast)
  5746     apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
  5754           apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
  5747     apply(simp)
  5755          apply(simp)
  5748     apply(simp)
  5756         apply(simp)
  5749     apply(assumption)
  5757        apply(assumption)
  5750     apply(rule allI)+
  5758       apply(rule allI)+
  5751     apply(rule impI)
  5759       apply(rule impI)
  5752     apply(simp add: psubst_nsubst[symmetric])
  5760       apply(simp add: psubst_nsubst[symmetric])
  5753     apply(rule BINDING_implies_CAND)
  5761       apply(rule BINDING_implies_CAND)
  5754     apply(unfold BINDINGc_def)
  5762       apply(unfold BINDINGc_def)
  5755     apply(simp)
  5763       apply(simp)
  5756     apply(rule_tac x="a" in exI)
  5764       apply(rule_tac x="a" in exI)
  5757     apply(rule_tac x="((x,ca,Q)#\<theta>_n),\<theta>_c<M>" in exI)
  5765       apply(rule_tac x="((x,ca,Q)#\<theta>_n),\<theta>_c<M>" in exI)
  5758     apply(simp)
  5766       apply(simp)
  5759     apply(rule allI)+
  5767       apply(rule allI)+
  5760     apply(rule impI)
  5768       apply(rule impI)
  5761     apply(rule_tac t="((x,ca,Q)#\<theta>_n),\<theta>_c<M>{a:=(xaa).Pa}" and 
  5769       apply(rule_tac t="((x,ca,Q)#\<theta>_n),\<theta>_c<M>{a:=(xaa).Pa}" and 
  5762                    s="((x,ca,Q)#\<theta>_n),((a,xaa,Pa)#\<theta>_c)<M>" in subst)
  5770         s="((x,ca,Q)#\<theta>_n),((a,xaa,Pa)#\<theta>_c)<M>" in subst)
  5763     apply(rule psubst_csubst)
  5771        apply(rule psubst_csubst)
  5764     apply(simp add: fresh_prod fresh_atm fresh_list_cons)
  5772        apply(simp add: fresh_prod fresh_atm fresh_list_cons)
  5765     apply(drule_tac x="(x,ca,Q)#\<theta>_n" in meta_spec)
  5773       apply(drule_tac x="(x,ca,Q)#\<theta>_n" in meta_spec)
  5766     apply(drule_tac x="(a,xaa,Pa)#\<theta>_c" in meta_spec)
  5774       apply(drule_tac x="(a,xaa,Pa)#\<theta>_c" in meta_spec)
  5767     apply(drule meta_mp)
  5775       apply(drule meta_mp)
  5768     apply(rule ncloses_extend)
  5776        apply(rule ncloses_extend)
  5769     apply(assumption)
  5777           apply(assumption)
  5770     apply(simp)
  5778          apply(simp)
  5771     apply(simp)
  5779         apply(simp)
  5772     apply(simp)
  5780        apply(simp)
  5773     apply(drule meta_mp)
  5781       apply(drule meta_mp)
  5774     apply(rule ccloses_extend)
  5782        apply(rule ccloses_extend)
  5775     apply(rule ccloses_subset)
  5783           apply(rule ccloses_subset)
  5776     apply(assumption)
  5784            apply(assumption)
  5777     apply(blast)
  5785           apply(blast)
  5778     apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
  5786          apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
  5779     apply(simp)
  5787         apply(simp)
  5780     apply(simp)
  5788        apply(simp)
  5781     apply(assumption)
  5789       apply(assumption)
  5782     apply(simp)
  5790      apply(simp)
  5783     apply(blast)
  5791     apply(blast)
  5784     done
  5792     done
  5785 next
  5793 next
  5786   case (TCut a \<Delta> N x \<Gamma> M B \<theta>_n \<theta>_c)
  5794   case (TCut a \<Delta> N x \<Gamma> M B \<theta>_n \<theta>_c)
  5787   then show ?case 
  5795   then show ?case 
  5788     apply -
  5796     apply -
  5789     apply(case_tac "\<forall>y. M\<noteq>Ax y a")
  5797     apply(case_tac "\<forall>y. M\<noteq>Ax y a")
  5790     apply(case_tac "\<forall>c. N\<noteq>Ax x c")
  5798      apply(case_tac "\<forall>c. N\<noteq>Ax x c")
  5791     apply(simp)
  5799       apply(simp)
       
  5800       apply(rule_tac B="B" in CUT_SNa)
       
  5801        apply(rule BINDING_implies_CAND)
       
  5802        apply(unfold BINDINGc_def)
       
  5803        apply(simp)
       
  5804        apply(rule_tac x="a" in exI)
       
  5805        apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
       
  5806        apply(simp)
       
  5807        apply(rule allI)
       
  5808        apply(rule allI)
       
  5809        apply(rule impI)
       
  5810        apply(simp add: psubst_csubst[symmetric]) (*?*)
       
  5811        apply(drule_tac x="\<theta>_n" in meta_spec)
       
  5812        apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
       
  5813        apply(drule meta_mp)
       
  5814         apply(assumption)
       
  5815        apply(drule meta_mp)
       
  5816         apply(rule ccloses_extend) 
       
  5817            apply(assumption)
       
  5818           apply(assumption)
       
  5819          apply(assumption)
       
  5820         apply(assumption)
       
  5821        apply(assumption)
       
  5822       apply(rule BINDING_implies_CAND)
       
  5823       apply(unfold BINDINGn_def)
       
  5824       apply(simp)
       
  5825       apply(rule_tac x="x" in exI)
       
  5826       apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
       
  5827       apply(simp)
       
  5828       apply(rule allI)
       
  5829       apply(rule allI)
       
  5830       apply(rule impI)
       
  5831       apply(simp add: psubst_nsubst[symmetric]) (*?*)
       
  5832       apply(rotate_tac 11)
       
  5833       apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
       
  5834       apply(drule_tac x="\<theta>_c" in meta_spec)
       
  5835       apply(drule meta_mp)
       
  5836        apply(rule ncloses_extend)
       
  5837           apply(assumption)
       
  5838          apply(assumption)
       
  5839         apply(assumption)
       
  5840        apply(assumption)
       
  5841       apply(drule_tac meta_mp)
       
  5842        apply(assumption)
       
  5843       apply(assumption)
       
  5844       (* cases at least one axiom *)
       
  5845      apply(simp (no_asm_use))
       
  5846      apply(erule exE)
       
  5847      apply(simp del: psubst.simps)
       
  5848      apply(drule typing_Ax_elim2)
       
  5849      apply(auto simp add: trm.inject)[1]
       
  5850      apply(rule_tac B="B" in CUT_SNa)
       
  5851       (* left term *)
       
  5852       apply(rule BINDING_implies_CAND)
       
  5853       apply(unfold BINDINGc_def)
       
  5854       apply(simp)
       
  5855       apply(rule_tac x="a" in exI)
       
  5856       apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
       
  5857       apply(simp)
       
  5858       apply(rule allI)+
       
  5859       apply(rule impI)
       
  5860       apply(drule_tac x="\<theta>_n" in meta_spec)
       
  5861       apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
       
  5862       apply(drule meta_mp)
       
  5863        apply(assumption)
       
  5864       apply(drule meta_mp)
       
  5865        apply(rule ccloses_extend) 
       
  5866           apply(assumption)
       
  5867          apply(assumption)
       
  5868         apply(assumption)
       
  5869        apply(assumption)
       
  5870       apply(simp add: psubst_csubst[symmetric]) (*?*)
       
  5871       (* right term -axiom *)
       
  5872      apply(drule ccloses_elim)
       
  5873       apply(assumption)
       
  5874      apply(erule exE)+
       
  5875      apply(erule conjE)
       
  5876      apply(frule_tac y="x" in lookupd_cmaps)
       
  5877      apply(drule cmaps_fresh)
       
  5878       apply(assumption)
       
  5879      apply(simp)
       
  5880      apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
       
  5881       apply(simp)
       
  5882      apply(simp add: ntrm.inject)
       
  5883      apply(simp add: alpha fresh_prod fresh_atm)
       
  5884      apply(rule sym)
       
  5885      apply(rule nrename_swap)
       
  5886      apply(simp)
       
  5887       (* M is axiom *)
       
  5888     apply(simp)
       
  5889     apply(auto)[1]
       
  5890       (* both are axioms *)
       
  5891      apply(rule_tac B="B" in CUT_SNa)
       
  5892       apply(drule typing_Ax_elim1)
       
  5893       apply(drule ncloses_elim)
       
  5894        apply(assumption)
       
  5895       apply(erule exE)+
       
  5896       apply(erule conjE)
       
  5897       apply(frule_tac a="a" in lookupc_nmaps)
       
  5898       apply(drule_tac a="a" in nmaps_fresh)
       
  5899        apply(assumption)
       
  5900       apply(simp)
       
  5901       apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
       
  5902        apply(simp)
       
  5903       apply(simp add: ctrm.inject)
       
  5904       apply(simp add: alpha fresh_prod fresh_atm)
       
  5905       apply(rule sym)
       
  5906       apply(rule crename_swap)
       
  5907       apply(simp)
       
  5908      apply(drule typing_Ax_elim2)
       
  5909      apply(drule ccloses_elim)
       
  5910       apply(assumption)
       
  5911      apply(erule exE)+
       
  5912      apply(erule conjE)
       
  5913      apply(frule_tac y="x" in lookupd_cmaps)
       
  5914      apply(drule cmaps_fresh)
       
  5915       apply(assumption)
       
  5916      apply(simp)
       
  5917      apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
       
  5918       apply(simp)
       
  5919      apply(simp add: ntrm.inject)
       
  5920      apply(simp add: alpha fresh_prod fresh_atm)
       
  5921      apply(rule sym)
       
  5922      apply(rule nrename_swap)
       
  5923      apply(simp)
       
  5924       (* N is not axioms *)
  5792     apply(rule_tac B="B" in CUT_SNa)
  5925     apply(rule_tac B="B" in CUT_SNa)
  5793     apply(rule BINDING_implies_CAND)
  5926       (* left term *)
  5794     apply(unfold BINDINGc_def)
  5927      apply(drule typing_Ax_elim1)
  5795     apply(simp)
  5928      apply(drule ncloses_elim)
  5796     apply(rule_tac x="a" in exI)
  5929       apply(assumption)
  5797     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
  5930      apply(erule exE)+
  5798     apply(simp)
  5931      apply(erule conjE)
  5799     apply(rule allI)
  5932      apply(frule_tac a="a" in lookupc_nmaps)
  5800     apply(rule allI)
  5933      apply(drule_tac a="a" in nmaps_fresh)
  5801     apply(rule impI)
  5934       apply(assumption)
  5802     apply(simp add: psubst_csubst[symmetric]) (*?*)
  5935      apply(simp)
  5803     apply(drule_tac x="\<theta>_n" in meta_spec)
  5936      apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
  5804     apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
  5937       apply(simp)
  5805     apply(drule meta_mp)
  5938      apply(simp add: ctrm.inject)
  5806     apply(assumption)
  5939      apply(simp add: alpha fresh_prod fresh_atm)
  5807     apply(drule meta_mp)
  5940      apply(rule sym)
  5808     apply(rule ccloses_extend) 
  5941      apply(rule crename_swap)
  5809     apply(assumption)
  5942      apply(simp)
  5810     apply(assumption)
       
  5811     apply(assumption)
       
  5812     apply(assumption)
       
  5813     apply(assumption)
       
  5814     apply(rule BINDING_implies_CAND)
       
  5815     apply(unfold BINDINGn_def)
       
  5816     apply(simp)
       
  5817     apply(rule_tac x="x" in exI)
       
  5818     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
       
  5819     apply(simp)
       
  5820     apply(rule allI)
       
  5821     apply(rule allI)
       
  5822     apply(rule impI)
       
  5823     apply(simp add: psubst_nsubst[symmetric]) (*?*)
       
  5824     apply(rotate_tac 11)
       
  5825     apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
       
  5826     apply(drule_tac x="\<theta>_c" in meta_spec)
       
  5827     apply(drule meta_mp)
       
  5828     apply(rule ncloses_extend)
       
  5829     apply(assumption)
       
  5830     apply(assumption)
       
  5831     apply(assumption)
       
  5832     apply(assumption)
       
  5833     apply(drule_tac meta_mp)
       
  5834     apply(assumption)
       
  5835     apply(assumption)
       
  5836     (* cases at least one axiom *)
       
  5837     apply(simp (no_asm_use))
       
  5838     apply(erule exE)
       
  5839     apply(simp del: psubst.simps)
       
  5840     apply(drule typing_Ax_elim2)
       
  5841     apply(auto simp add: trm.inject)[1]
       
  5842     apply(rule_tac B="B" in CUT_SNa)
       
  5843     (* left term *)
       
  5844     apply(rule BINDING_implies_CAND)
       
  5845     apply(unfold BINDINGc_def)
       
  5846     apply(simp)
       
  5847     apply(rule_tac x="a" in exI)
       
  5848     apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
       
  5849     apply(simp)
       
  5850     apply(rule allI)+
       
  5851     apply(rule impI)
       
  5852     apply(drule_tac x="\<theta>_n" in meta_spec)
       
  5853     apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
       
  5854     apply(drule meta_mp)
       
  5855     apply(assumption)
       
  5856     apply(drule meta_mp)
       
  5857     apply(rule ccloses_extend) 
       
  5858     apply(assumption)
       
  5859     apply(assumption)
       
  5860     apply(assumption)
       
  5861     apply(assumption)
       
  5862     apply(simp add: psubst_csubst[symmetric]) (*?*)
       
  5863     (* right term -axiom *)
       
  5864     apply(drule ccloses_elim)
       
  5865     apply(assumption)
       
  5866     apply(erule exE)+
       
  5867     apply(erule conjE)
       
  5868     apply(frule_tac y="x" in lookupd_cmaps)
       
  5869     apply(drule cmaps_fresh)
       
  5870     apply(assumption)
       
  5871     apply(simp)
       
  5872     apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
       
  5873     apply(simp)
       
  5874     apply(simp add: ntrm.inject)
       
  5875     apply(simp add: alpha fresh_prod fresh_atm)
       
  5876     apply(rule sym)
       
  5877     apply(rule nrename_swap)
       
  5878     apply(simp)
       
  5879     (* M is axiom *)
       
  5880     apply(simp)
       
  5881     apply(auto)[1]
       
  5882     (* both are axioms *)
       
  5883     apply(rule_tac B="B" in CUT_SNa)
       
  5884     apply(drule typing_Ax_elim1)
       
  5885     apply(drule ncloses_elim)
       
  5886     apply(assumption)
       
  5887     apply(erule exE)+
       
  5888     apply(erule conjE)
       
  5889     apply(frule_tac a="a" in lookupc_nmaps)
       
  5890     apply(drule_tac a="a" in nmaps_fresh)
       
  5891     apply(assumption)
       
  5892     apply(simp)
       
  5893     apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
       
  5894     apply(simp)
       
  5895     apply(simp add: ctrm.inject)
       
  5896     apply(simp add: alpha fresh_prod fresh_atm)
       
  5897     apply(rule sym)
       
  5898     apply(rule crename_swap)
       
  5899     apply(simp)
       
  5900     apply(drule typing_Ax_elim2)
       
  5901     apply(drule ccloses_elim)
       
  5902     apply(assumption)
       
  5903     apply(erule exE)+
       
  5904     apply(erule conjE)
       
  5905     apply(frule_tac y="x" in lookupd_cmaps)
       
  5906     apply(drule cmaps_fresh)
       
  5907     apply(assumption)
       
  5908     apply(simp)
       
  5909     apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
       
  5910     apply(simp)
       
  5911     apply(simp add: ntrm.inject)
       
  5912     apply(simp add: alpha fresh_prod fresh_atm)
       
  5913     apply(rule sym)
       
  5914     apply(rule nrename_swap)
       
  5915     apply(simp)
       
  5916     (* N is not axioms *)
       
  5917     apply(rule_tac B="B" in CUT_SNa)
       
  5918     (* left term *)
       
  5919     apply(drule typing_Ax_elim1)
       
  5920     apply(drule ncloses_elim)
       
  5921     apply(assumption)
       
  5922     apply(erule exE)+
       
  5923     apply(erule conjE)
       
  5924     apply(frule_tac a="a" in lookupc_nmaps)
       
  5925     apply(drule_tac a="a" in nmaps_fresh)
       
  5926     apply(assumption)
       
  5927     apply(simp)
       
  5928     apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
       
  5929     apply(simp)
       
  5930     apply(simp add: ctrm.inject)
       
  5931     apply(simp add: alpha fresh_prod fresh_atm)
       
  5932     apply(rule sym)
       
  5933     apply(rule crename_swap)
       
  5934     apply(simp)
       
  5935     apply(rule BINDING_implies_CAND)
  5943     apply(rule BINDING_implies_CAND)
  5936     apply(unfold BINDINGn_def)
  5944     apply(unfold BINDINGn_def)
  5937     apply(simp)
  5945     apply(simp)
  5938     apply(rule_tac x="x" in exI)
  5946     apply(rule_tac x="x" in exI)
  5939     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5947     apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
  5944     apply(simp add: psubst_nsubst[symmetric]) (*?*)
  5952     apply(simp add: psubst_nsubst[symmetric]) (*?*)
  5945     apply(rotate_tac 10)
  5953     apply(rotate_tac 10)
  5946     apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
  5954     apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
  5947     apply(drule_tac x="\<theta>_c" in meta_spec)
  5955     apply(drule_tac x="\<theta>_c" in meta_spec)
  5948     apply(drule meta_mp)
  5956     apply(drule meta_mp)
  5949     apply(rule ncloses_extend)
  5957      apply(rule ncloses_extend)
  5950     apply(assumption)
  5958         apply(assumption)
  5951     apply(assumption)
  5959        apply(assumption)
  5952     apply(assumption)
  5960       apply(assumption)
  5953     apply(assumption)
  5961      apply(assumption)
  5954     apply(drule_tac meta_mp)
  5962     apply(drule_tac meta_mp)
  5955     apply(assumption)
  5963      apply(assumption)
  5956     apply(assumption)
  5964     apply(assumption)
  5957     done
  5965     done
  5958 qed
  5966 qed
  5959 
  5967 
  5960 primrec "idn" :: "(name\<times>ty) list\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list" where
  5968 primrec "idn" :: "(name\<times>ty) list\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list" where
  5965   "idc [] x    = []"
  5973   "idc [] x    = []"
  5966 | "idc (p#\<Delta>) x = ((fst p),x,Ax x (fst p))#(idc \<Delta> x)"
  5974 | "idc (p#\<Delta>) x = ((fst p),x,Ax x (fst p))#(idc \<Delta> x)"
  5967 
  5975 
  5968 lemma idn_eqvt[eqvt]:
  5976 lemma idn_eqvt[eqvt]:
  5969   fixes pi1::"name prm"
  5977   fixes pi1::"name prm"
  5970   and   pi2::"coname prm"
  5978     and   pi2::"coname prm"
  5971   shows "(pi1\<bullet>(idn \<Gamma> a)) = idn (pi1\<bullet>\<Gamma>) (pi1\<bullet>a)"
  5979   shows "(pi1\<bullet>(idn \<Gamma> a)) = idn (pi1\<bullet>\<Gamma>) (pi1\<bullet>a)"
  5972   and   "(pi2\<bullet>(idn \<Gamma> a)) = idn (pi2\<bullet>\<Gamma>) (pi2\<bullet>a)"
  5980     and   "(pi2\<bullet>(idn \<Gamma> a)) = idn (pi2\<bullet>\<Gamma>) (pi2\<bullet>a)"
  5973 apply(induct \<Gamma>)
  5981    apply(induct \<Gamma>)
  5974 apply(auto)
  5982      apply(auto)
  5975 done
  5983   done
  5976 
  5984 
  5977 lemma idc_eqvt[eqvt]:
  5985 lemma idc_eqvt[eqvt]:
  5978   fixes pi1::"name prm"
  5986   fixes pi1::"name prm"
  5979   and   pi2::"coname prm"
  5987     and   pi2::"coname prm"
  5980   shows "(pi1\<bullet>(idc \<Delta> x)) = idc (pi1\<bullet>\<Delta>) (pi1\<bullet>x)"
  5988   shows "(pi1\<bullet>(idc \<Delta> x)) = idc (pi1\<bullet>\<Delta>) (pi1\<bullet>x)"
  5981   and   "(pi2\<bullet>(idc \<Delta> x)) = idc (pi2\<bullet>\<Delta>) (pi2\<bullet>x)"
  5989     and   "(pi2\<bullet>(idc \<Delta> x)) = idc (pi2\<bullet>\<Delta>) (pi2\<bullet>x)"
  5982 apply(induct \<Delta>)
  5990    apply(induct \<Delta>)
  5983 apply(auto)
  5991      apply(auto)
  5984 done
  5992   done
  5985 
  5993 
  5986 lemma ccloses_id:
  5994 lemma ccloses_id:
  5987   shows "(idc \<Delta> x) ccloses \<Delta>"
  5995   shows "(idc \<Delta> x) ccloses \<Delta>"
  5988 apply(induct \<Delta>)
  5996   apply(induct \<Delta>)
  5989 apply(auto simp add: ccloses_def)
  5997    apply(auto simp add: ccloses_def)
  5990 apply(rule Ax_in_CANDs)
  5998    apply(rule Ax_in_CANDs)
  5991 apply(rule Ax_in_CANDs)
  5999   apply(rule Ax_in_CANDs)
  5992 done
  6000   done
  5993 
  6001 
  5994 lemma ncloses_id:
  6002 lemma ncloses_id:
  5995   shows "(idn \<Gamma> a) ncloses \<Gamma>"
  6003   shows "(idn \<Gamma> a) ncloses \<Gamma>"
  5996 apply(induct \<Gamma>)
  6004   apply(induct \<Gamma>)
  5997 apply(auto simp add: ncloses_def)
  6005    apply(auto simp add: ncloses_def)
  5998 apply(rule Ax_in_CANDs)
  6006    apply(rule Ax_in_CANDs)
  5999 apply(rule Ax_in_CANDs)
  6007   apply(rule Ax_in_CANDs)
  6000 done
  6008   done
  6001 
  6009 
  6002 lemma fresh_idn:
  6010 lemma fresh_idn:
  6003   fixes x::"name"
  6011   fixes x::"name"
  6004   and   a::"coname"
  6012     and   a::"coname"
  6005   shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<sharp>idn \<Gamma> a"
  6013   shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<sharp>idn \<Gamma> a"
  6006   and   "a\<sharp>(\<Gamma>,b) \<Longrightarrow> a\<sharp>idn \<Gamma> b"
  6014     and   "a\<sharp>(\<Gamma>,b) \<Longrightarrow> a\<sharp>idn \<Gamma> b"
  6007 apply(induct \<Gamma>)
  6015    apply(induct \<Gamma>)
  6008 apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
  6016      apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
  6009 done
  6017   done
  6010 
  6018 
  6011 lemma fresh_idc:
  6019 lemma fresh_idc:
  6012   fixes x::"name"
  6020   fixes x::"name"
  6013   and   a::"coname"
  6021     and   a::"coname"
  6014   shows "x\<sharp>(\<Delta>,y) \<Longrightarrow> x\<sharp>idc \<Delta> y"
  6022   shows "x\<sharp>(\<Delta>,y) \<Longrightarrow> x\<sharp>idc \<Delta> y"
  6015   and   "a\<sharp>\<Delta>  \<Longrightarrow> a\<sharp>idc \<Delta> y"
  6023     and   "a\<sharp>\<Delta>  \<Longrightarrow> a\<sharp>idc \<Delta> y"
  6016 apply(induct \<Delta>)
  6024    apply(induct \<Delta>)
  6017 apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
  6025      apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
  6018 done
  6026   done
  6019 
  6027 
  6020 lemma idc_cmaps:
  6028 lemma idc_cmaps:
  6021   assumes a: "idc \<Delta> y cmaps b to Some (x,M)"
  6029   assumes a: "idc \<Delta> y cmaps b to Some (x,M)"
  6022   shows "M=Ax x b"
  6030   shows "M=Ax x b"
  6023 using a
  6031   using a
  6024 apply(induct \<Delta>)
  6032   apply(induct \<Delta>)
  6025 apply(auto)
  6033    apply(auto)
  6026 apply(case_tac "b=a")
  6034   apply(case_tac "b=a")
  6027 apply(auto)
  6035    apply(auto)
  6028 done
  6036   done
  6029 
  6037 
  6030 lemma idn_nmaps:
  6038 lemma idn_nmaps:
  6031   assumes a: "idn \<Gamma> a nmaps x to Some (b,M)"
  6039   assumes a: "idn \<Gamma> a nmaps x to Some (b,M)"
  6032   shows "M=Ax x b"
  6040   shows "M=Ax x b"
  6033 using a
  6041   using a
  6034 apply(induct \<Gamma>)
  6042   apply(induct \<Gamma>)
  6035 apply(auto)
  6043    apply(auto)
  6036 apply(case_tac "aa=x")
  6044   apply(case_tac "aa=x")
  6037 apply(auto)
  6045    apply(auto)
  6038 done
  6046   done
  6039 
  6047 
  6040 lemma lookup1:
  6048 lemma lookup1:
  6041   assumes a: "x\<sharp>(idn \<Gamma> b)"
  6049   assumes a: "x\<sharp>(idn \<Gamma> b)"
  6042   shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupa x a \<theta>_c"
  6050   shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupa x a \<theta>_c"
  6043 using a
  6051   using a
  6044 apply(induct \<Gamma>)
  6052   apply(induct \<Gamma>)
  6045 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  6053    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  6046 done
  6054   done
  6047 
  6055 
  6048 lemma lookup2:
  6056 lemma lookup2:
  6049   assumes a: "\<not>(x\<sharp>(idn \<Gamma> b))"
  6057   assumes a: "\<not>(x\<sharp>(idn \<Gamma> b))"
  6050   shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupb x a \<theta>_c b (Ax x b)"
  6058   shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupb x a \<theta>_c b (Ax x b)"
  6051 using a
  6059   using a
  6052 apply(induct \<Gamma>)
  6060   apply(induct \<Gamma>)
  6053 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6061    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6054 done
  6062   done
  6055 
  6063 
  6056 lemma lookup3:
  6064 lemma lookup3:
  6057   assumes a: "a\<sharp>(idc \<Delta> y)"
  6065   assumes a: "a\<sharp>(idc \<Delta> y)"
  6058   shows "lookupa x a (idc \<Delta> y) = Ax x a"
  6066   shows "lookupa x a (idc \<Delta> y) = Ax x a"
  6059 using a
  6067   using a
  6060 apply(induct \<Delta>)
  6068   apply(induct \<Delta>)
  6061 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  6069    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
  6062 done
  6070   done
  6063 
  6071 
  6064 lemma lookup4:
  6072 lemma lookup4:
  6065   assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
  6073   assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
  6066   shows "lookupa x a (idc \<Delta> y) = Cut <a>.(Ax x a) (y).Ax y a"
  6074   shows "lookupa x a (idc \<Delta> y) = Cut <a>.(Ax x a) (y).Ax y a"
  6067 using a
  6075   using a
  6068 apply(induct \<Delta>)
  6076   apply(induct \<Delta>)
  6069 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6077    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6070 done
  6078   done
  6071 
  6079 
  6072 lemma lookup5:
  6080 lemma lookup5:
  6073   assumes a: "a\<sharp>(idc \<Delta> y)"
  6081   assumes a: "a\<sharp>(idc \<Delta> y)"
  6074   shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (x).Ax x a"
  6082   shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (x).Ax x a"
  6075 using a
  6083   using a
  6076 apply(induct \<Delta>)
  6084   apply(induct \<Delta>)
  6077 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6085    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6078 done
  6086   done
  6079 
  6087 
  6080 lemma lookup6:
  6088 lemma lookup6:
  6081   assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
  6089   assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
  6082   shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (y).Ax y a"
  6090   shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (y).Ax y a"
  6083 using a
  6091   using a
  6084 apply(induct \<Delta>)
  6092   apply(induct \<Delta>)
  6085 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6093    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6086 done
  6094   done
  6087 
  6095 
  6088 lemma lookup7:
  6096 lemma lookup7:
  6089   shows "lookupc x a (idn \<Gamma> b) = Ax x a"
  6097   shows "lookupc x a (idn \<Gamma> b) = Ax x a"
  6090 apply(induct \<Gamma>)
  6098   apply(induct \<Gamma>)
  6091 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6099    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6092 done
  6100   done
  6093 
  6101 
  6094 lemma lookup8:
  6102 lemma lookup8:
  6095   shows "lookupd x a (idc \<Delta> y) = Ax x a"
  6103   shows "lookupd x a (idc \<Delta> y) = Ax x a"
  6096 apply(induct \<Delta>)
  6104   apply(induct \<Delta>)
  6097 apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6105    apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
  6098 done
  6106   done
  6099 
  6107 
  6100 lemma id_redu:
  6108 lemma id_redu:
  6101   shows "(idn \<Gamma> x),(idc \<Delta> a)<M> \<longrightarrow>\<^sub>a* M"
  6109   shows "(idn \<Gamma> x),(idc \<Delta> a)<M> \<longrightarrow>\<^sub>a* M"
  6102 apply(nominal_induct M avoiding: \<Gamma> \<Delta> x a rule: trm.strong_induct)
  6110   apply(nominal_induct M avoiding: \<Gamma> \<Delta> x a rule: trm.strong_induct)
  6103 apply(auto)
  6111              apply(auto)
  6104 (* Ax *)
  6112     (* Ax *)
  6105 apply(case_tac "name\<sharp>(idn \<Gamma> x)")
  6113              apply(case_tac "name\<sharp>(idn \<Gamma> x)")
  6106 apply(simp add: lookup1)
  6114               apply(simp add: lookup1)
  6107 apply(case_tac "coname\<sharp>(idc \<Delta> a)")
  6115               apply(case_tac "coname\<sharp>(idc \<Delta> a)")
  6108 apply(simp add: lookup3)
  6116                apply(simp add: lookup3)
  6109 apply(simp add: lookup4)
  6117               apply(simp add: lookup4)
  6110 apply(rule a_star_trans)
  6118               apply(rule a_star_trans)
  6111 apply(rule a_starI)
  6119                apply(rule a_starI)
  6112 apply(rule al_redu)
  6120                apply(rule al_redu)
  6113 apply(rule better_LAxR_intro)
  6121                apply(rule better_LAxR_intro)
  6114 apply(rule fic.intros)
  6122                apply(rule fic.intros)
  6115 apply(simp)
  6123               apply(simp)
  6116 apply(simp add: lookup2)
  6124              apply(simp add: lookup2)
  6117 apply(case_tac "coname\<sharp>(idc \<Delta> a)")
  6125              apply(case_tac "coname\<sharp>(idc \<Delta> a)")
  6118 apply(simp add: lookup5)
  6126               apply(simp add: lookup5)
  6119 apply(rule a_star_trans)
  6127               apply(rule a_star_trans)
  6120 apply(rule a_starI)
  6128                apply(rule a_starI)
  6121 apply(rule al_redu)
  6129                apply(rule al_redu)
  6122 apply(rule better_LAxR_intro)
  6130                apply(rule better_LAxR_intro)
  6123 apply(rule fic.intros)
  6131                apply(rule fic.intros)
  6124 apply(simp)
  6132               apply(simp)
  6125 apply(simp add: lookup6)
  6133              apply(simp add: lookup6)
  6126 apply(rule a_star_trans)
  6134              apply(rule a_star_trans)
  6127 apply(rule a_starI)
  6135               apply(rule a_starI)
  6128 apply(rule al_redu)
  6136               apply(rule al_redu)
  6129 apply(rule better_LAxR_intro)
  6137               apply(rule better_LAxR_intro)
  6130 apply(rule fic.intros)
  6138               apply(rule fic.intros)
  6131 apply(simp)
  6139              apply(simp)
  6132 (* Cut *)
  6140     (* Cut *)
  6133 apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name psubst_fresh_coname fresh_atm fresh_prod )[1]
  6141             apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name psubst_fresh_coname fresh_atm fresh_prod )[1]
  6134 apply(simp add: lookup7 lookup8)
  6142                apply(simp add: lookup7 lookup8)
  6135 apply(simp add: lookup7 lookup8)
  6143               apply(simp add: lookup7 lookup8)
  6136 apply(simp add: a_star_Cut)
  6144               apply(simp add: a_star_Cut)
  6137 apply(simp add: lookup7 lookup8)
  6145              apply(simp add: lookup7 lookup8)
  6138 apply(simp add: a_star_Cut)
  6146              apply(simp add: a_star_Cut)
  6139 apply(simp add: a_star_Cut)
  6147             apply(simp add: a_star_Cut)
  6140 (* NotR *)
  6148     (* NotR *)
  6141 apply(simp add: fresh_idn fresh_idc)
  6149            apply(simp add: fresh_idn fresh_idc)
  6142 apply(case_tac "findc (idc \<Delta> a) coname")
  6150            apply(case_tac "findc (idc \<Delta> a) coname")
  6143 apply(simp)
  6151             apply(simp)
  6144 apply(simp add: a_star_NotR)
  6152             apply(simp add: a_star_NotR)
  6145 apply(auto)[1]
  6153            apply(auto)[1]
  6146 apply(generate_fresh "coname")
  6154            apply(generate_fresh "coname")
  6147 apply(fresh_fun_simp)
  6155            apply(fresh_fun_simp)
  6148 apply(drule idc_cmaps)
  6156            apply(drule idc_cmaps)
  6149 apply(simp)
  6157            apply(simp)
  6150 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6158            apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6151 apply(rule a_star_trans)
  6159             apply(rule a_star_trans)
  6152 apply(rule a_starI)
  6160              apply(rule a_starI)
  6153 apply(rule al_redu)
  6161              apply(rule al_redu)
  6154 apply(rule better_LAxR_intro)
  6162              apply(rule better_LAxR_intro)
  6155 apply(rule fic.intros)
  6163              apply(rule fic.intros)
  6156 apply(assumption)
  6164              apply(assumption)
  6157 apply(simp add: crename_fresh)
  6165             apply(simp add: crename_fresh)
  6158 apply(simp add: a_star_NotR)
  6166             apply(simp add: a_star_NotR)
  6159 apply(rule psubst_fresh_coname)
  6167            apply(rule psubst_fresh_coname)
  6160 apply(rule fresh_idn)
  6168              apply(rule fresh_idn)
  6161 apply(simp)
  6169              apply(simp)
  6162 apply(rule fresh_idc)
  6170             apply(rule fresh_idc)
  6163 apply(simp)
  6171             apply(simp)
  6164 apply(simp)
  6172            apply(simp)
  6165 (* NotL *)
  6173     (* NotL *)
  6166 apply(simp add: fresh_idn fresh_idc)
  6174           apply(simp add: fresh_idn fresh_idc)
  6167 apply(case_tac "findn (idn \<Gamma> x) name")
  6175           apply(case_tac "findn (idn \<Gamma> x) name")
  6168 apply(simp)
  6176            apply(simp)
  6169 apply(simp add: a_star_NotL)
  6177            apply(simp add: a_star_NotL)
  6170 apply(auto)[1]
  6178           apply(auto)[1]
  6171 apply(generate_fresh "name")
  6179           apply(generate_fresh "name")
  6172 apply(fresh_fun_simp)
  6180           apply(fresh_fun_simp)
  6173 apply(drule idn_nmaps)
  6181           apply(drule idn_nmaps)
  6174 apply(simp)
  6182           apply(simp)
  6175 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6183           apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6176 apply(rule a_star_trans)
  6184            apply(rule a_star_trans)
  6177 apply(rule a_starI)
  6185             apply(rule a_starI)
  6178 apply(rule al_redu)
  6186             apply(rule al_redu)
  6179 apply(rule better_LAxL_intro)
  6187             apply(rule better_LAxL_intro)
  6180 apply(rule fin.intros)
  6188             apply(rule fin.intros)
  6181 apply(assumption)
  6189             apply(assumption)
  6182 apply(simp add: nrename_fresh)
  6190            apply(simp add: nrename_fresh)
  6183 apply(simp add: a_star_NotL)
  6191            apply(simp add: a_star_NotL)
  6184 apply(rule psubst_fresh_name)
  6192           apply(rule psubst_fresh_name)
  6185 apply(rule fresh_idn)
  6193             apply(rule fresh_idn)
  6186 apply(simp)
  6194             apply(simp)
  6187 apply(rule fresh_idc)
  6195            apply(rule fresh_idc)
  6188 apply(simp)
  6196            apply(simp)
  6189 apply(simp)
  6197           apply(simp)
  6190 (* AndR *)
  6198     (* AndR *)
  6191 apply(simp add: fresh_idn fresh_idc)
  6199          apply(simp add: fresh_idn fresh_idc)
  6192 apply(case_tac "findc (idc \<Delta> a) coname3")
  6200          apply(case_tac "findc (idc \<Delta> a) coname3")
  6193 apply(simp)
  6201           apply(simp)
  6194 apply(simp add: a_star_AndR)
  6202           apply(simp add: a_star_AndR)
  6195 apply(auto)[1]
  6203          apply(auto)[1]
  6196 apply(generate_fresh "coname")
  6204          apply(generate_fresh "coname")
  6197 apply(fresh_fun_simp)
  6205          apply(fresh_fun_simp)
  6198 apply(drule idc_cmaps)
  6206          apply(drule idc_cmaps)
  6199 apply(simp)
  6207          apply(simp)
  6200 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
  6208          apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
  6201 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
  6209           apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
  6202 apply(rule a_star_trans)
  6210            apply(rule a_star_trans)
  6203 apply(rule a_starI)
  6211             apply(rule a_starI)
  6204 apply(rule al_redu)
  6212             apply(rule al_redu)
  6205 apply(rule better_LAxR_intro)
  6213             apply(rule better_LAxR_intro)
  6206 apply(rule fic.intros)
  6214             apply(rule fic.intros)
  6207 apply(simp add: abs_fresh)
  6215              apply(simp add: abs_fresh)
  6208 apply(simp add: abs_fresh)
  6216             apply(simp add: abs_fresh)
  6209 apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name crename_fresh fresh_atm fresh_prod )[1]
  6217            apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name crename_fresh fresh_atm fresh_prod )[1]
  6210 apply(rule aux3)
  6218            apply(rule aux3)
  6211 apply(rule crename.simps)
  6219             apply(rule crename.simps)
  6212 apply(auto simp add: fresh_prod fresh_atm)[1]
  6220               apply(auto simp add: fresh_prod fresh_atm)[1]
  6213 apply(rule psubst_fresh_coname)
  6221               apply(rule psubst_fresh_coname)
  6214 apply(rule fresh_idn)
  6222                 apply(rule fresh_idn)
  6215 apply(simp add: fresh_prod fresh_atm)
  6223                 apply(simp add: fresh_prod fresh_atm)
  6216 apply(rule fresh_idc)
  6224                apply(rule fresh_idc)
  6217 apply(simp)
  6225                apply(simp)
  6218 apply(simp)
  6226               apply(simp)
  6219 apply(auto simp add: fresh_prod fresh_atm)[1]
  6227              apply(auto simp add: fresh_prod fresh_atm)[1]
  6220 apply(rule psubst_fresh_coname)
  6228              apply(rule psubst_fresh_coname)
  6221 apply(rule fresh_idn)
  6229                apply(rule fresh_idn)
  6222 apply(simp add: fresh_prod fresh_atm)
  6230                apply(simp add: fresh_prod fresh_atm)
  6223 apply(rule fresh_idc)
  6231               apply(rule fresh_idc)
  6224 apply(simp)
  6232               apply(simp)
  6225 apply(simp)
  6233              apply(simp)
  6226 apply(simp)
  6234             apply(simp)
  6227 apply(simp)
  6235            apply(simp)
  6228 apply(simp add: crename_fresh)
  6236            apply(simp add: crename_fresh)
  6229 apply(simp add: a_star_AndR)
  6237            apply(simp add: a_star_AndR)
  6230 apply(rule psubst_fresh_coname)
  6238           apply(rule psubst_fresh_coname)
  6231 apply(rule fresh_idn)
  6239             apply(rule fresh_idn)
  6232 apply(simp)
  6240             apply(simp)
  6233 apply(rule fresh_idc)
  6241            apply(rule fresh_idc)
  6234 apply(simp)
  6242            apply(simp)
  6235 apply(simp)
  6243           apply(simp)
  6236 apply(rule psubst_fresh_coname)
  6244          apply(rule psubst_fresh_coname)
  6237 apply(rule fresh_idn)
  6245            apply(rule fresh_idn)
  6238 apply(simp)
  6246            apply(simp)
  6239 apply(rule fresh_idc)
  6247           apply(rule fresh_idc)
  6240 apply(simp)
  6248           apply(simp)
  6241 apply(simp)
  6249          apply(simp)
  6242 (* AndL1 *)
  6250     (* AndL1 *)
  6243 apply(simp add: fresh_idn fresh_idc)
  6251         apply(simp add: fresh_idn fresh_idc)
  6244 apply(case_tac "findn (idn \<Gamma> x) name2")
  6252         apply(case_tac "findn (idn \<Gamma> x) name2")
  6245 apply(simp)
  6253          apply(simp)
  6246 apply(simp add: a_star_AndL1)
  6254          apply(simp add: a_star_AndL1)
  6247 apply(auto)[1]
  6255         apply(auto)[1]
  6248 apply(generate_fresh "name")
  6256         apply(generate_fresh "name")
  6249 apply(fresh_fun_simp)
  6257         apply(fresh_fun_simp)
  6250 apply(drule idn_nmaps)
  6258         apply(drule idn_nmaps)
  6251 apply(simp)
  6259         apply(simp)
  6252 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6260         apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6253 apply(rule a_star_trans)
  6261          apply(rule a_star_trans)
  6254 apply(rule a_starI)
  6262           apply(rule a_starI)
  6255 apply(rule al_redu)
  6263           apply(rule al_redu)
  6256 apply(rule better_LAxL_intro)
  6264           apply(rule better_LAxL_intro)
  6257 apply(rule fin.intros)
  6265           apply(rule fin.intros)
  6258 apply(simp add: abs_fresh)
  6266           apply(simp add: abs_fresh)
  6259 apply(rule aux3)
  6267          apply(rule aux3)
  6260 apply(rule nrename.simps)
  6268           apply(rule nrename.simps)
  6261 apply(auto simp add: fresh_prod fresh_atm)[1]
  6269           apply(auto simp add: fresh_prod fresh_atm)[1]
  6262 apply(simp)
  6270          apply(simp)
  6263 apply(simp add: nrename_fresh)
  6271          apply(simp add: nrename_fresh)
  6264 apply(simp add: a_star_AndL1)
  6272          apply(simp add: a_star_AndL1)
  6265 apply(rule psubst_fresh_name)
  6273         apply(rule psubst_fresh_name)
  6266 apply(rule fresh_idn)
  6274           apply(rule fresh_idn)
  6267 apply(simp)
  6275           apply(simp)
  6268 apply(rule fresh_idc)
  6276          apply(rule fresh_idc)
  6269 apply(simp)
  6277          apply(simp)
  6270 apply(simp)
  6278         apply(simp)
  6271 (* AndL2 *)
  6279     (* AndL2 *)
  6272 apply(simp add: fresh_idn fresh_idc)
  6280        apply(simp add: fresh_idn fresh_idc)
  6273 apply(case_tac "findn (idn \<Gamma> x) name2")
  6281        apply(case_tac "findn (idn \<Gamma> x) name2")
  6274 apply(simp)
  6282         apply(simp)
  6275 apply(simp add: a_star_AndL2)
  6283         apply(simp add: a_star_AndL2)
  6276 apply(auto)[1]
  6284        apply(auto)[1]
  6277 apply(generate_fresh "name")
  6285        apply(generate_fresh "name")
  6278 apply(fresh_fun_simp)
  6286        apply(fresh_fun_simp)
  6279 apply(drule idn_nmaps)
  6287        apply(drule idn_nmaps)
  6280 apply(simp)
  6288        apply(simp)
  6281 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6289        apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6282 apply(rule a_star_trans)
  6290         apply(rule a_star_trans)
  6283 apply(rule a_starI)
  6291          apply(rule a_starI)
  6284 apply(rule al_redu)
  6292          apply(rule al_redu)
  6285 apply(rule better_LAxL_intro)
  6293          apply(rule better_LAxL_intro)
  6286 apply(rule fin.intros)
  6294          apply(rule fin.intros)
  6287 apply(simp add: abs_fresh)
  6295          apply(simp add: abs_fresh)
  6288 apply(rule aux3)
  6296         apply(rule aux3)
  6289 apply(rule nrename.simps)
  6297          apply(rule nrename.simps)
  6290 apply(auto simp add: fresh_prod fresh_atm)[1]
  6298          apply(auto simp add: fresh_prod fresh_atm)[1]
  6291 apply(simp)
  6299         apply(simp)
  6292 apply(simp add: nrename_fresh)
  6300         apply(simp add: nrename_fresh)
  6293 apply(simp add: a_star_AndL2)
  6301         apply(simp add: a_star_AndL2)
  6294 apply(rule psubst_fresh_name)
  6302        apply(rule psubst_fresh_name)
  6295 apply(rule fresh_idn)
  6303          apply(rule fresh_idn)
  6296 apply(simp)
  6304          apply(simp)
  6297 apply(rule fresh_idc)
  6305         apply(rule fresh_idc)
  6298 apply(simp)
  6306         apply(simp)
  6299 apply(simp)
  6307        apply(simp)
  6300 (* OrR1 *)
  6308     (* OrR1 *)
  6301 apply(simp add: fresh_idn fresh_idc)
  6309       apply(simp add: fresh_idn fresh_idc)
  6302 apply(case_tac "findc (idc \<Delta> a) coname2")
  6310       apply(case_tac "findc (idc \<Delta> a) coname2")
  6303 apply(simp)
  6311        apply(simp)
  6304 apply(simp add: a_star_OrR1)
  6312        apply(simp add: a_star_OrR1)
  6305 apply(auto)[1]
  6313       apply(auto)[1]
  6306 apply(generate_fresh "coname")
  6314       apply(generate_fresh "coname")
  6307 apply(fresh_fun_simp)
  6315       apply(fresh_fun_simp)
  6308 apply(drule idc_cmaps)
  6316       apply(drule idc_cmaps)
  6309 apply(simp)
  6317       apply(simp)
  6310 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6318       apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6311 apply(rule a_star_trans)
  6319        apply(rule a_star_trans)
  6312 apply(rule a_starI)
  6320         apply(rule a_starI)
  6313 apply(rule al_redu)
  6321         apply(rule al_redu)
  6314 apply(rule better_LAxR_intro)
  6322         apply(rule better_LAxR_intro)
  6315 apply(rule fic.intros)
  6323         apply(rule fic.intros)
  6316 apply(simp add: abs_fresh)
  6324         apply(simp add: abs_fresh)
  6317 apply(rule aux3)
  6325        apply(rule aux3)
  6318 apply(rule crename.simps)
  6326         apply(rule crename.simps)
  6319 apply(auto simp add: fresh_prod fresh_atm)[1]
  6327         apply(auto simp add: fresh_prod fresh_atm)[1]
  6320 apply(simp)
  6328        apply(simp)
  6321 apply(simp add: crename_fresh)
  6329        apply(simp add: crename_fresh)
  6322 apply(simp add: a_star_OrR1)
  6330        apply(simp add: a_star_OrR1)
  6323 apply(rule psubst_fresh_coname)
  6331       apply(rule psubst_fresh_coname)
  6324 apply(rule fresh_idn)
  6332         apply(rule fresh_idn)
  6325 apply(simp)
  6333         apply(simp)
  6326 apply(rule fresh_idc)
  6334        apply(rule fresh_idc)
  6327 apply(simp)
  6335        apply(simp)
  6328 apply(simp)
  6336       apply(simp)
  6329 (* OrR2 *)
  6337     (* OrR2 *)
  6330 apply(simp add: fresh_idn fresh_idc)
  6338      apply(simp add: fresh_idn fresh_idc)
  6331 apply(case_tac "findc (idc \<Delta> a) coname2")
  6339      apply(case_tac "findc (idc \<Delta> a) coname2")
  6332 apply(simp)
  6340       apply(simp)
  6333 apply(simp add: a_star_OrR2)
  6341       apply(simp add: a_star_OrR2)
  6334 apply(auto)[1]
  6342      apply(auto)[1]
  6335 apply(generate_fresh "coname")
  6343      apply(generate_fresh "coname")
  6336 apply(fresh_fun_simp)
  6344      apply(fresh_fun_simp)
  6337 apply(drule idc_cmaps)
  6345      apply(drule idc_cmaps)
  6338 apply(simp)
  6346      apply(simp)
  6339 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6347      apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6340 apply(rule a_star_trans)
  6348       apply(rule a_star_trans)
  6341 apply(rule a_starI)
  6349        apply(rule a_starI)
  6342 apply(rule al_redu)
  6350        apply(rule al_redu)
  6343 apply(rule better_LAxR_intro)
  6351        apply(rule better_LAxR_intro)
  6344 apply(rule fic.intros)
  6352        apply(rule fic.intros)
  6345 apply(simp add: abs_fresh)
  6353        apply(simp add: abs_fresh)
  6346 apply(rule aux3)
  6354       apply(rule aux3)
  6347 apply(rule crename.simps)
  6355        apply(rule crename.simps)
  6348 apply(auto simp add: fresh_prod fresh_atm)[1]
  6356        apply(auto simp add: fresh_prod fresh_atm)[1]
  6349 apply(simp)
  6357       apply(simp)
  6350 apply(simp add: crename_fresh)
  6358       apply(simp add: crename_fresh)
  6351 apply(simp add: a_star_OrR2)
  6359       apply(simp add: a_star_OrR2)
  6352 apply(rule psubst_fresh_coname)
  6360      apply(rule psubst_fresh_coname)
  6353 apply(rule fresh_idn)
  6361        apply(rule fresh_idn)
  6354 apply(simp)
  6362        apply(simp)
  6355 apply(rule fresh_idc)
  6363       apply(rule fresh_idc)
  6356 apply(simp)
  6364       apply(simp)
  6357 apply(simp)
  6365      apply(simp)
  6358 (* OrL *)
  6366     (* OrL *)
  6359 apply(simp add: fresh_idn fresh_idc)
  6367     apply(simp add: fresh_idn fresh_idc)
  6360 apply(case_tac "findn (idn \<Gamma> x) name3")
  6368     apply(case_tac "findn (idn \<Gamma> x) name3")
  6361 apply(simp)
  6369      apply(simp)
  6362 apply(simp add: a_star_OrL)
  6370      apply(simp add: a_star_OrL)
  6363 apply(auto)[1]
  6371     apply(auto)[1]
  6364 apply(generate_fresh "name")
  6372     apply(generate_fresh "name")
  6365 apply(fresh_fun_simp)
  6373     apply(fresh_fun_simp)
  6366 apply(drule idn_nmaps)
  6374     apply(drule idn_nmaps)
  6367 apply(simp)
  6375     apply(simp)
  6368 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
  6376     apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
  6369 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
  6377      apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
  6370 apply(rule a_star_trans)
  6378       apply(rule a_star_trans)
  6371 apply(rule a_starI)
  6379        apply(rule a_starI)
  6372 apply(rule al_redu)
  6380        apply(rule al_redu)
  6373 apply(rule better_LAxL_intro)
  6381        apply(rule better_LAxL_intro)
  6374 apply(rule fin.intros)
  6382        apply(rule fin.intros)
  6375 apply(simp add: abs_fresh)
  6383         apply(simp add: abs_fresh)
  6376 apply(simp add: abs_fresh)
  6384        apply(simp add: abs_fresh)
  6377 apply(rule aux3)
  6385       apply(rule aux3)
  6378 apply(rule nrename.simps)
  6386        apply(rule nrename.simps)
  6379 apply(auto simp add: fresh_prod fresh_atm)[1]
  6387          apply(auto simp add: fresh_prod fresh_atm)[1]
  6380 apply(rule psubst_fresh_name)
  6388          apply(rule psubst_fresh_name)
  6381 apply(rule fresh_idn)
  6389            apply(rule fresh_idn)
  6382 apply(simp)
  6390            apply(simp)
  6383 apply(rule fresh_idc)
  6391           apply(rule fresh_idc)
  6384 apply(simp add: fresh_prod fresh_atm)
  6392           apply(simp add: fresh_prod fresh_atm)
  6385 apply(simp)
  6393          apply(simp)
  6386 apply(auto simp add: fresh_prod fresh_atm)[1]
  6394         apply(auto simp add: fresh_prod fresh_atm)[1]
  6387 apply(rule psubst_fresh_name)
  6395         apply(rule psubst_fresh_name)
  6388 apply(rule fresh_idn)
  6396           apply(rule fresh_idn)
  6389 apply(simp)
  6397           apply(simp)
  6390 apply(rule fresh_idc)
  6398          apply(rule fresh_idc)
  6391 apply(simp add: fresh_prod fresh_atm)
  6399          apply(simp add: fresh_prod fresh_atm)
  6392 apply(simp)
  6400         apply(simp)
  6393 apply(simp)
  6401        apply(simp)
  6394 apply(simp)
  6402       apply(simp)
  6395 apply(simp add: nrename_fresh)
  6403       apply(simp add: nrename_fresh)
  6396 apply(simp add: a_star_OrL)
  6404       apply(simp add: a_star_OrL)
  6397 apply(rule psubst_fresh_name)
  6405      apply(rule psubst_fresh_name)
  6398 apply(rule fresh_idn)
  6406        apply(rule fresh_idn)
  6399 apply(simp)
  6407        apply(simp)
  6400 apply(rule fresh_idc)
  6408       apply(rule fresh_idc)
  6401 apply(simp)
  6409       apply(simp)
  6402 apply(simp)
  6410      apply(simp)
  6403 apply(rule psubst_fresh_name)
  6411     apply(rule psubst_fresh_name)
  6404 apply(rule fresh_idn)
  6412       apply(rule fresh_idn)
  6405 apply(simp)
  6413       apply(simp)
  6406 apply(rule fresh_idc)
  6414      apply(rule fresh_idc)
  6407 apply(simp)
  6415      apply(simp)
  6408 apply(simp)
  6416     apply(simp)
  6409 (* ImpR *)
  6417     (* ImpR *)
  6410 apply(simp add: fresh_idn fresh_idc)
  6418    apply(simp add: fresh_idn fresh_idc)
  6411 apply(case_tac "findc (idc \<Delta> a) coname2")
  6419    apply(case_tac "findc (idc \<Delta> a) coname2")
  6412 apply(simp)
  6420     apply(simp)
  6413 apply(simp add: a_star_ImpR)
  6421     apply(simp add: a_star_ImpR)
  6414 apply(auto)[1]
  6422    apply(auto)[1]
  6415 apply(generate_fresh "coname")
  6423    apply(generate_fresh "coname")
  6416 apply(fresh_fun_simp)
  6424    apply(fresh_fun_simp)
  6417 apply(drule idc_cmaps)
  6425    apply(drule idc_cmaps)
  6418 apply(simp)
  6426    apply(simp)
  6419 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6427    apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
  6420 apply(rule a_star_trans)
  6428     apply(rule a_star_trans)
  6421 apply(rule a_starI)
  6429      apply(rule a_starI)
  6422 apply(rule al_redu)
  6430      apply(rule al_redu)
  6423 apply(rule better_LAxR_intro)
  6431      apply(rule better_LAxR_intro)
  6424 apply(rule fic.intros)
  6432      apply(rule fic.intros)
  6425 apply(simp add: abs_fresh)
  6433      apply(simp add: abs_fresh)
  6426 apply(rule aux3)
  6434     apply(rule aux3)
  6427 apply(rule crename.simps)
  6435      apply(rule crename.simps)
  6428 apply(auto simp add: fresh_prod fresh_atm)[1]
  6436      apply(auto simp add: fresh_prod fresh_atm)[1]
  6429 apply(simp)
  6437     apply(simp)
  6430 apply(simp add: crename_fresh)
  6438     apply(simp add: crename_fresh)
  6431 apply(simp add: a_star_ImpR)
  6439     apply(simp add: a_star_ImpR)
  6432 apply(rule psubst_fresh_coname)
  6440    apply(rule psubst_fresh_coname)
  6433 apply(rule fresh_idn)
  6441      apply(rule fresh_idn)
  6434 apply(simp)
  6442      apply(simp)
  6435 apply(rule fresh_idc)
  6443     apply(rule fresh_idc)
  6436 apply(simp)
  6444     apply(simp)
  6437 apply(simp)
  6445    apply(simp)
  6438 (* ImpL *)
  6446     (* ImpL *)
  6439 apply(simp add: fresh_idn fresh_idc)
  6447   apply(simp add: fresh_idn fresh_idc)
  6440 apply(case_tac "findn (idn \<Gamma> x) name2")
  6448   apply(case_tac "findn (idn \<Gamma> x) name2")
  6441 apply(simp)
  6449    apply(simp)
  6442 apply(simp add: a_star_ImpL)
  6450    apply(simp add: a_star_ImpL)
  6443 apply(auto)[1]
  6451   apply(auto)[1]
  6444 apply(generate_fresh "name")
  6452   apply(generate_fresh "name")
  6445 apply(fresh_fun_simp)
  6453   apply(fresh_fun_simp)
  6446 apply(drule idn_nmaps)
  6454   apply(drule idn_nmaps)
  6447 apply(simp)
  6455   apply(simp)
  6448 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
  6456   apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
  6449 apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
  6457    apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
  6450 apply(rule a_star_trans)
  6458     apply(rule a_star_trans)
  6451 apply(rule a_starI)
  6459      apply(rule a_starI)
  6452 apply(rule al_redu)
  6460      apply(rule al_redu)
  6453 apply(rule better_LAxL_intro)
  6461      apply(rule better_LAxL_intro)
  6454 apply(rule fin.intros)
  6462      apply(rule fin.intros)
  6455 apply(simp add: abs_fresh)
  6463       apply(simp add: abs_fresh)
  6456 apply(simp add: abs_fresh)
  6464      apply(simp add: abs_fresh)
  6457 apply(rule aux3)
  6465     apply(rule aux3)
  6458 apply(rule nrename.simps)
  6466      apply(rule nrename.simps)
  6459 apply(auto simp add: fresh_prod fresh_atm)[1]
  6467       apply(auto simp add: fresh_prod fresh_atm)[1]
  6460 apply(rule psubst_fresh_coname)
  6468       apply(rule psubst_fresh_coname)
  6461 apply(rule fresh_idn)
  6469         apply(rule fresh_idn)
  6462 apply(simp add: fresh_atm)
  6470         apply(simp add: fresh_atm)
  6463 apply(rule fresh_idc)
  6471        apply(rule fresh_idc)
  6464 apply(simp add: fresh_prod fresh_atm)
  6472        apply(simp add: fresh_prod fresh_atm)
  6465 apply(simp)
  6473       apply(simp)
  6466 apply(auto simp add: fresh_prod fresh_atm)[1]
  6474      apply(auto simp add: fresh_prod fresh_atm)[1]
  6467 apply(rule psubst_fresh_name)
  6475      apply(rule psubst_fresh_name)
  6468 apply(rule fresh_idn)
  6476        apply(rule fresh_idn)
  6469 apply(simp)
  6477        apply(simp)
  6470 apply(rule fresh_idc)
  6478       apply(rule fresh_idc)
  6471 apply(simp add: fresh_prod fresh_atm)
  6479       apply(simp add: fresh_prod fresh_atm)
  6472 apply(simp)
  6480      apply(simp)
  6473 apply(simp)
  6481     apply(simp)
  6474 apply(simp add: nrename_fresh)
  6482     apply(simp add: nrename_fresh)
  6475 apply(simp add: a_star_ImpL)
  6483     apply(simp add: a_star_ImpL)
  6476 apply(rule psubst_fresh_name)
  6484    apply(rule psubst_fresh_name)
  6477 apply(rule fresh_idn)
  6485      apply(rule fresh_idn)
  6478 apply(simp)
  6486      apply(simp)
  6479 apply(rule fresh_idc)
  6487     apply(rule fresh_idc)
  6480 apply(simp)
  6488     apply(simp)
  6481 apply(simp)
  6489    apply(simp)
  6482 apply(rule psubst_fresh_name)
  6490   apply(rule psubst_fresh_name)
  6483 apply(rule fresh_idn)
  6491     apply(rule fresh_idn)
  6484 apply(simp)
  6492     apply(simp)
  6485 apply(rule fresh_idc)
  6493    apply(rule fresh_idc)
  6486 apply(simp)
  6494    apply(simp)
  6487 apply(simp)
  6495   apply(simp)
  6488 done
  6496   done
  6489 
  6497 
  6490 theorem ALL_SNa:
  6498 theorem ALL_SNa:
  6491   assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
  6499   assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
  6492   shows "SNa M"
  6500   shows "SNa M"
  6493 proof -
  6501 proof -