src/HOL/Nominal/Examples/Class3.thy
author blanchet
Wed, 12 Feb 2014 08:37:06 +0100
changeset 55417 01fbfb60c33e
parent 55413 a8e96847523c
child 61943 7fba644ed827
permissions -rw-r--r--
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems * * * more transition of 'xxx_rec' to 'rec_xxx' and same for case * * * compile * * * 'rename_tac's to avoid referring to generated names * * * more robust scripts with 'rename_tac' * * * 'where' -> 'of' * * * 'where' -> 'of' * * * renamed 'xxx_rec' to 'rec_xxx'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     1
theory Class3
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     2
imports Class2
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     3
begin
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     4
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     5
text {* 3rd Main Lemma *}
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     6
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     7
lemma Cut_a_redu_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
     8
  assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
     9
  shows "(\<exists>M'. R = Cut <a>.M' (x).N \<and> M \<longrightarrow>\<^sub>a M') \<or>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
    10
         (\<exists>N'. R = Cut <a>.M (x).N' \<and> N \<longrightarrow>\<^sub>a N') \<or>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
    11
         (Cut <a>.M (x).N \<longrightarrow>\<^sub>c R) \<or>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
    12
         (Cut <a>.M (x).N \<longrightarrow>\<^sub>l R)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    13
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    14
apply(erule_tac a_redu.cases)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    15
apply(simp_all)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    16
apply(simp_all add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    17
apply(rule disjI1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    18
apply(auto simp add: alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    19
apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    20
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    21
apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    22
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    23
apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    24
apply(rule disjI1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    25
apply(auto simp add: alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    26
apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    27
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    28
apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    29
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    30
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    31
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    32
lemma Cut_c_redu_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
    33
  assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>c R"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    34
  shows "(R = M{a:=(x).N} \<and> \<not>fic M a) \<or>
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    35
         (R = N{x:=<a>.M} \<and> \<not>fin N x)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    36
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    37
apply(erule_tac c_redu.cases)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    38
apply(simp_all)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    39
apply(simp_all add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    40
apply(rule disjI1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    41
apply(auto simp add: alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    42
apply(simp add: subst_rename fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    43
apply(simp add: subst_rename fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    44
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    45
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    46
apply(simp add: subst_rename fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    47
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    48
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    49
apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    50
apply(auto simp add: alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    51
apply(simp add: subst_rename fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    52
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    53
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    54
apply(simp add: subst_rename fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    55
apply(simp add: subst_rename fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    56
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    57
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    58
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    59
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    60
lemma not_fic_crename_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    61
  assumes a: "fic M c" "c\<sharp>(a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    62
  shows "fic (M[a\<turnstile>c>b]) c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    63
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    64
apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    65
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    66
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    67
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    68
lemma not_fic_crename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    69
  assumes a: "\<not>(fic (M[a\<turnstile>c>b]) c)" "c\<sharp>(a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    70
  shows "\<not>(fic M c)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    71
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    72
apply(auto dest:  not_fic_crename_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    73
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    74
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    75
lemma not_fin_crename_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    76
  assumes a: "fin M y"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    77
  shows "fin (M[a\<turnstile>c>b]) y" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    78
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    79
apply(nominal_induct M avoiding: a b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    80
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    81
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    82
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    83
lemma not_fin_crename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    84
  assumes a: "\<not>(fin (M[a\<turnstile>c>b]) y)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    85
  shows "\<not>(fin M y)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    86
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    87
apply(auto dest:  not_fin_crename_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    88
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    89
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    90
lemma crename_fresh_interesting1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    91
  fixes c::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    92
  assumes a: "c\<sharp>(M[a\<turnstile>c>b])" "c\<sharp>(a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    93
  shows "c\<sharp>M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    94
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    95
apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    96
apply(auto split: if_splits simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    97
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    98
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    99
lemma crename_fresh_interesting2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   100
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   101
  assumes a: "x\<sharp>(M[a\<turnstile>c>b])" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   102
  shows "x\<sharp>M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   103
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   104
apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   105
apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   106
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   107
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   108
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   109
lemma fic_crename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   110
  assumes a: "fic (M[a\<turnstile>c>b]) c" "c\<sharp>(a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   111
  shows "fic M c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   112
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   113
apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   114
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   115
           split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   116
apply(auto dest: crename_fresh_interesting1 simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   117
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   118
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   119
lemma fin_crename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   120
  assumes a: "fin (M[a\<turnstile>c>b]) x"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   121
  shows "fin M x" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   122
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   123
apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   124
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   125
           split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   126
apply(auto dest: crename_fresh_interesting2 simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   127
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   128
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   129
lemma crename_Cut:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   130
  assumes a: "R[a\<turnstile>c>b] = Cut <c>.M (x).N" "c\<sharp>(a,b,N,R)" "x\<sharp>(M,R)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   132
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   133
apply(nominal_induct R avoiding: a b c x M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   134
apply(auto split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   135
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   136
apply(auto simp add: alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   137
apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   138
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   139
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   140
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   141
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   142
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   143
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   144
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   145
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   146
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   147
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   148
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   149
apply(auto simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   150
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   151
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   152
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   153
apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   154
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   155
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   156
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   157
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   158
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   159
apply(auto simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   160
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   161
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   162
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   163
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   164
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   165
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   166
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   167
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   168
lemma crename_NotR:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   169
  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>(a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   170
  shows "\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   171
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   172
apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   173
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   174
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   175
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   176
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   177
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   178
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   179
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   180
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   181
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   182
lemma crename_NotR':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   183
  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   185
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   186
apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   187
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   188
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   189
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   190
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   191
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   192
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   193
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   194
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   195
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   196
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   197
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   198
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   199
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   200
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   201
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   202
lemma crename_NotR_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   203
  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   204
  shows "(a=c \<and> a=b) \<or> (a\<noteq>c)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   205
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   206
apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   207
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   208
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   209
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   210
lemma crename_NotL:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   211
  assumes a: "R[a\<turnstile>c>b] = NotL <c>.N y" "c\<sharp>(R,a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   212
  shows "\<exists>N'. (R = NotL <c>.N' y) \<and> N'[a\<turnstile>c>b] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   213
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   214
apply(nominal_induct R avoiding: a b c y N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   215
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   216
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   217
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   218
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   219
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   220
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   221
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   222
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   223
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   224
lemma crename_AndL1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   225
  assumes a: "R[a\<turnstile>c>b] = AndL1 (x).N y" "x\<sharp>R"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   226
  shows "\<exists>N'. (R = AndL1 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   227
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   228
apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   229
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   230
apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   231
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   232
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   233
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   234
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   235
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   236
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   237
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   238
lemma crename_AndL2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   239
  assumes a: "R[a\<turnstile>c>b] = AndL2 (x).N y" "x\<sharp>R"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   240
  shows "\<exists>N'. (R = AndL2 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   241
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   242
apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   243
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   244
apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   245
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   246
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   247
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   248
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   249
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   250
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   251
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   252
lemma crename_AndR_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   253
  assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   254
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   255
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   256
apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   257
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   258
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   259
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   260
lemma crename_AndR:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   263
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   264
apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   265
apply(auto split: if_splits simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   266
apply(simp add: fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   267
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   268
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   269
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   270
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   271
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   272
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   273
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   274
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   275
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   276
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   277
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   278
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   279
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   280
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   281
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   282
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   283
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   284
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   285
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   286
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   287
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   288
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   289
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   290
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   291
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   292
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   293
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   294
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   295
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   296
lemma crename_AndR':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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>
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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')"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   300
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   301
apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   302
apply(auto split: if_splits simp add: trm.inject alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   303
apply(auto split: if_splits simp add: trm.inject alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   304
apply(auto split: if_splits simp add: trm.inject alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   305
apply(auto split: if_splits simp add: trm.inject alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   306
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   307
apply(case_tac "coname3=a")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   308
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   309
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   310
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   311
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   312
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   313
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   314
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   315
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   316
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   317
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   318
apply(drule_tac s="trm2[a\<turnstile>c>e]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   319
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   320
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   321
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   322
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   323
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   324
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   325
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   326
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   327
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   328
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   329
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   330
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   331
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   332
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   333
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   334
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   335
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   336
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   337
lemma crename_OrR1_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   338
  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.M e" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   339
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   340
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   341
apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   342
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   343
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   344
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   345
lemma crename_OrR1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   346
  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   347
  shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   348
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   349
apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   350
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   351
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   352
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   353
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   354
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   355
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   356
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   357
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   358
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   359
lemma crename_OrR1':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   360
  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   361
  shows "(\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   362
         (\<exists>N'. (R = OrR1 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   363
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   364
apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   365
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   366
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   367
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   368
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   369
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   370
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   371
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   372
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   373
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   374
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   375
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   376
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   377
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   378
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   379
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   380
lemma crename_OrR2_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   381
  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.M e" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   382
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   383
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   384
apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   385
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   386
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   387
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   388
lemma crename_OrR2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   389
  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   390
  shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   391
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   392
apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   393
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   394
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   395
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   396
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   397
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   398
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   399
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   400
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   401
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   402
lemma crename_OrR2':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   403
  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   404
  shows "(\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   405
         (\<exists>N'. (R = OrR2 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   406
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   407
apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   408
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   409
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   410
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   411
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   412
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   413
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   414
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   415
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   416
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   417
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   418
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   419
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   420
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   421
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   422
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   423
lemma crename_OrL:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   426
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   427
apply(nominal_induct R avoiding: a b x y z M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   428
apply(auto split: if_splits simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   429
apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   430
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   431
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   432
apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   433
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   434
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   435
apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   436
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   437
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   438
apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   439
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   440
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   441
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   442
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   443
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   444
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   445
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   446
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   447
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   448
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   449
lemma crename_ImpL:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   452
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   453
apply(nominal_induct R avoiding: a b c y z M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   454
apply(auto split: if_splits simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   455
apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   456
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   457
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   458
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   459
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   460
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   461
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   462
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   463
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   464
apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   465
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   466
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   467
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   468
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   469
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   470
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   471
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   472
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   473
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   474
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   475
lemma crename_ImpR_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   476
  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.M e" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   477
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   478
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   479
apply(nominal_induct R avoiding: x a b c e M rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   480
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   481
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   482
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   483
lemma crename_ImpR:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   485
  shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   486
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   487
apply(nominal_induct R avoiding: a b x c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   488
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   489
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   490
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   491
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   492
apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   493
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   494
apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   495
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   496
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   497
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   498
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   499
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   500
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   501
lemma crename_ImpR':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   502
  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "x\<sharp>R" "d\<sharp>a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   503
  shows "(\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   504
         (\<exists>N'. (R = ImpR (x).<c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   505
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   506
apply(nominal_induct R avoiding: x a b c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   507
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject abs_perm calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   508
apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   509
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   510
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   511
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   512
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   513
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   514
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   515
apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   516
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   517
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   518
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   519
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   520
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   521
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   522
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   523
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   524
lemma crename_ax2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   525
  assumes a: "N[a\<turnstile>c>b] = Ax x c"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   526
  shows "\<exists>d. N = Ax x d"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   527
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   528
apply(nominal_induct N avoiding: a b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   529
apply(auto split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   530
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   531
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   532
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   533
lemma crename_interesting1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   534
  assumes a: "distinct [a,b,c]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   535
  shows "M[a\<turnstile>c>c][c\<turnstile>c>b] = M[c\<turnstile>c>b][a\<turnstile>c>b]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   536
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   537
apply(nominal_induct M avoiding: a c b rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   538
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   539
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   540
apply(rotate_tac 12)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   541
apply(drule_tac x="a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   542
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   543
apply(drule_tac x="c" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   544
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   545
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   546
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   547
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   548
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   549
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   550
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   551
lemma crename_interesting2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   552
  assumes a: "a\<noteq>c" "a\<noteq>d" "a\<noteq>b" "c\<noteq>d" "b\<noteq>c"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   553
  shows "M[a\<turnstile>c>b][c\<turnstile>c>d] = M[c\<turnstile>c>d][a\<turnstile>c>b]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   554
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   555
apply(nominal_induct M avoiding: a c b d rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   556
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   557
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   558
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   559
lemma crename_interesting3:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   560
  shows "M[a\<turnstile>c>c][x\<turnstile>n>y] = M[x\<turnstile>n>y][a\<turnstile>c>c]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   561
apply(nominal_induct M avoiding: a c x y rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   562
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   563
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   564
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   565
lemma crename_credu:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
   566
  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>c M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
   567
  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>c M0"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   568
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   569
apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: c_redu.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   570
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   571
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   572
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   573
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   574
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   575
apply(rule_tac x="M'{a:=(x).N'}" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   576
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   577
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   578
apply(rule c_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   579
apply(auto dest: not_fic_crename)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   580
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   581
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   582
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   583
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   584
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   585
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   586
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   587
apply(rule_tac x="N'{x:=<a>.M'}" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   588
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   589
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   590
apply(rule c_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   591
apply(auto dest: not_fin_crename)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   592
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   593
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   594
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   595
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   596
lemma crename_lredu:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
   597
  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>l M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
   598
  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>l M0"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   599
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   600
apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: l_redu.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   601
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   602
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   603
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   604
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   605
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   606
apply(case_tac "aa=ba")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   607
apply(simp add: crename_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   608
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   609
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   610
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   611
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   612
apply(frule crename_ax2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   613
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   614
apply(case_tac "d=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   615
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   616
apply(rule_tac x="M'[a\<turnstile>c>aa]" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   617
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   618
apply(rule crename_interesting1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   619
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   620
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   621
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   622
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   623
apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   624
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   625
apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   626
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   627
apply(rule crename_interesting2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   628
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   629
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   630
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   631
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   632
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   633
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   634
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   635
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   636
apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   637
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   638
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   639
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   640
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   641
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   642
apply(case_tac "aa=b")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   643
apply(simp add: crename_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   644
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   645
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   646
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   647
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   648
apply(frule crename_ax2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   649
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   650
apply(case_tac "d=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   651
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   652
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   653
apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   654
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   655
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   656
apply(rule crename_interesting3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   657
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   658
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   659
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   660
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   661
(* LNot *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   662
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   663
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   664
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   665
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   666
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   667
apply(drule crename_NotR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   668
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   669
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   670
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   671
apply(drule crename_NotL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   672
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   673
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   674
apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   675
apply(simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   676
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   677
apply(auto simp add: fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   678
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   679
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   680
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   681
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   682
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   683
(* LAnd1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   684
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   685
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   686
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   687
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   688
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   689
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   690
apply(drule crename_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   691
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   692
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   693
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   694
apply(drule crename_AndL1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   695
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   696
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   697
apply(rule_tac x="Cut <a1>.M'a (x).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   698
apply(simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   699
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   700
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   701
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   702
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   703
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   704
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   705
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   706
(* LAnd2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   707
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   708
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   709
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   710
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   711
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   712
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   713
apply(drule crename_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   714
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   715
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   716
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   717
apply(drule crename_AndL2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   718
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   719
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   720
apply(rule_tac x="Cut <a2>.N'b (x).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   721
apply(simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   722
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   723
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   724
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   725
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   726
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   727
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   728
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   729
(* LOr1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   730
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   731
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   732
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   733
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   734
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   735
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   736
apply(drule crename_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   737
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   738
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   739
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   740
apply(drule crename_OrR1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   741
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   742
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   743
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   744
apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   745
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   746
apply(simp add: abs_fresh fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   747
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   748
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   749
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   750
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   751
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   752
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   753
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   754
(* LOr2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   755
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   756
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   757
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   758
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   759
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   760
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   761
apply(drule crename_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   762
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   763
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   764
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   765
apply(drule crename_OrR2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   766
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   767
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   768
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   769
apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   770
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   771
apply(simp add: abs_fresh fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   772
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   773
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   774
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   775
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   776
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   777
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   778
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   779
(* ImpL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   780
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   781
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   782
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   783
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   784
apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   785
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   786
apply(drule crename_ImpL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   787
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   788
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   789
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   790
apply(drule crename_ImpR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   791
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   792
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   793
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   794
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   795
apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   796
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   797
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   798
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   799
apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   800
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   801
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   802
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   803
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   804
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   805
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   806
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   807
lemma crename_aredu:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
   808
  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>a M'" "a\<noteq>b"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
   809
  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>a M0"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   810
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   811
apply(nominal_induct "M[a\<turnstile>c>b]" M' avoiding: M a b rule: a_redu.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   812
apply(drule  crename_lredu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   813
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   814
apply(drule  crename_credu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   815
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   816
(* Cut *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   817
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   818
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   819
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   820
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   821
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   822
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   823
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   824
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   825
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   826
apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   827
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   828
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   829
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   830
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   831
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   832
apply(drule crename_fresh_interesting2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   833
apply(simp add: fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   834
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   835
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   836
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   837
apply(drule crename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   838
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   839
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   840
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   841
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   842
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   843
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   844
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   845
apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   846
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   847
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   848
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   849
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   850
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   851
apply(drule crename_fresh_interesting1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   852
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   853
apply(simp add: fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   854
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   855
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   856
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   857
(* NotL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   858
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   859
apply(drule crename_NotL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   860
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   861
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   862
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   863
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   864
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   865
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   866
apply(rule_tac x="NotL <a>.M0 x" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   867
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   868
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   869
(* NotR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   870
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   871
apply(frule crename_NotR_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   872
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   873
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   874
apply(drule crename_NotR')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   875
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   876
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   877
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   878
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   879
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   880
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   881
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   882
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   883
apply(rule_tac x="NotR (x).M0 a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   884
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   885
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   886
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   887
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   888
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   889
apply(drule_tac x="a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   890
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   891
apply(rule_tac x="NotR (x).M0 aa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   892
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   893
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   894
(* AndR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   895
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   896
apply(frule crename_AndR_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   897
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   898
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   899
apply(drule crename_AndR')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   900
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   901
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   902
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   903
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   904
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   905
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   906
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   907
apply(drule_tac x="ba" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   908
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   909
apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   910
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   911
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   912
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   913
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   914
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   915
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   916
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   917
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   918
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   919
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   920
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   921
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   922
apply(drule_tac x="c" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   923
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   924
apply(rule_tac x="AndR <a>.M0 <b>.N' aa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   925
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   926
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   927
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   928
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   929
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   930
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   931
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   932
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   933
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   934
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   935
apply(frule crename_AndR_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   936
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   937
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   938
apply(drule crename_AndR')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   939
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   940
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   941
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   942
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   943
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   944
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   945
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   946
apply(drule_tac x="ba" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   947
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   948
apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   949
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   950
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   951
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   952
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   953
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   954
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   955
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   956
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   957
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   958
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   959
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   960
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   961
apply(drule_tac x="c" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   962
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   963
apply(rule_tac x="AndR <a>.M' <b>.M0 aa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   964
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   965
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   966
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   967
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   968
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   969
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   970
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   971
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   972
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   973
(* AndL1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   974
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   975
apply(drule crename_AndL1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   976
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   977
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   978
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   979
apply(drule_tac x="a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   980
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   981
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   982
apply(rule_tac x="AndL1 (x).M0 y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   983
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   984
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   985
(* AndL2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   986
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   987
apply(drule crename_AndL2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   988
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   989
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   990
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   991
apply(drule_tac x="a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   992
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   993
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   994
apply(rule_tac x="AndL2 (x).M0 y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   995
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   996
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   997
(* OrL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   998
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   999
apply(drule crename_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1000
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1001
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1002
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1003
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1004
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1005
apply(drule_tac x="a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1006
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1007
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1008
apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1009
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1010
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1011
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1012
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1013
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1014
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1015
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1016
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1017
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1018
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1019
apply(drule crename_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1020
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1021
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1022
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1023
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1024
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1025
apply(drule_tac x="a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1026
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1027
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1028
apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1029
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1030
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1031
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1032
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1033
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1034
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1035
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1036
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1037
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1038
(* OrR1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1039
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1040
apply(frule crename_OrR1_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1041
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1042
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1043
apply(drule crename_OrR1')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1044
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1045
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1046
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1047
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1048
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1049
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1050
apply(drule_tac x="ba" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1051
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1052
apply(rule_tac x="OrR1 <a>.M0 b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1053
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1054
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1055
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1056
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1057
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1058
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1059
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1060
apply(rule_tac x="OrR1 <a>.M0 aa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1061
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1062
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1063
(* OrR2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1064
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1065
apply(frule crename_OrR2_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1066
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1067
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1068
apply(drule crename_OrR2')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1069
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1070
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1071
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1072
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1073
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1074
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1075
apply(drule_tac x="ba" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1076
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1077
apply(rule_tac x="OrR2 <a>.M0 b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1078
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1079
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1080
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1081
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1082
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1083
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1084
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1085
apply(rule_tac x="OrR2 <a>.M0 aa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1086
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1087
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1088
(* ImpL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1089
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1090
apply(drule crename_ImpL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1091
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1092
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1093
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1094
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1095
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1096
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1097
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1098
apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1099
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1100
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1101
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1102
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1103
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1104
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1105
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1106
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1107
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1108
apply(drule crename_ImpL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1109
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1110
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1111
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1112
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1113
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1114
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1115
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1116
apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1117
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1118
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1119
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1120
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1121
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1122
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1123
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1124
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1125
(* ImpR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1126
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1127
apply(frule crename_ImpR_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1128
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1129
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1130
apply(drule crename_ImpR')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1131
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1132
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1133
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1134
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1135
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1136
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1137
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1138
apply(drule_tac x="ba" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1139
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1140
apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1141
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1142
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1143
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1144
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1145
apply(drule_tac x="aa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1146
apply(drule_tac x="b" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1147
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1148
apply(rule_tac x="ImpR (x).<a>.M0 aa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1149
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1150
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1151
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1152
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1153
lemma SNa_preserved_renaming1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1154
  assumes a: "SNa M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1155
  shows "SNa (M[a\<turnstile>c>b])"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1156
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1157
apply(induct rule: SNa_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1158
apply(case_tac "a=b")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1159
apply(simp add: crename_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1160
apply(rule SNaI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1161
apply(drule crename_aredu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1162
apply(blast)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1163
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1164
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1165
lemma nrename_interesting1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1166
  assumes a: "distinct [x,y,z]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1167
  shows "M[x\<turnstile>n>z][z\<turnstile>n>y] = M[z\<turnstile>n>y][x\<turnstile>n>y]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1168
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1169
apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1170
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1171
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1172
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1173
apply(rotate_tac 12)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1174
apply(drule_tac x="x" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1175
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1176
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1177
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1178
apply(drule_tac x="z" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1179
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1180
apply(rotate_tac 11)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1181
apply(drule_tac x="x" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1182
apply(rotate_tac 14)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1183
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1184
apply(rotate_tac 14)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1185
apply(drule_tac x="z" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1186
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1187
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1188
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1189
lemma nrename_interesting2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1190
  assumes a: "x\<noteq>z" "x\<noteq>u" "x\<noteq>y" "z\<noteq>u" "y\<noteq>z"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1191
  shows "M[x\<turnstile>n>y][z\<turnstile>n>u] = M[z\<turnstile>n>u][x\<turnstile>n>y]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1192
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1193
apply(nominal_induct M avoiding: x y z u rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1194
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1195
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1196
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1197
lemma not_fic_nrename_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1198
  assumes a: "fic M c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1199
  shows "fic (M[x\<turnstile>n>y]) c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1200
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1201
apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1202
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1203
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1204
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1205
lemma not_fic_nrename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1206
  assumes a: "\<not>(fic (M[x\<turnstile>n>y]) c)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1207
  shows "\<not>(fic M c)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1208
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1209
apply(auto dest:  not_fic_nrename_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1210
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1211
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1212
lemma fin_nrename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1213
  assumes a: "fin M z" "z\<sharp>(x,y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1214
  shows "fin (M[x\<turnstile>n>y]) z" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1215
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1216
apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1217
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1218
           split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1219
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1220
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1221
lemma nrename_fresh_interesting1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1222
  fixes z::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1223
  assumes a: "z\<sharp>(M[x\<turnstile>n>y])" "z\<sharp>(x,y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1224
  shows "z\<sharp>M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1225
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1226
apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1227
apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1228
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1229
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1230
lemma nrename_fresh_interesting2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1231
  fixes c::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1232
  assumes a: "c\<sharp>(M[x\<turnstile>n>y])" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1233
  shows "c\<sharp>M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1234
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1235
apply(nominal_induct M avoiding: x y c rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1236
apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1237
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1238
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1239
lemma fin_nrename2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1240
  assumes a: "fin (M[x\<turnstile>n>y]) z" "z\<sharp>(x,y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1241
  shows "fin M z" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1242
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1243
apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1244
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1245
           split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1246
apply(auto dest: nrename_fresh_interesting1 simp add: fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1247
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1248
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1249
lemma nrename_Cut:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1250
  assumes a: "R[x\<turnstile>n>y] = Cut <c>.M (z).N" "c\<sharp>(N,R)" "z\<sharp>(x,y,M,R)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1252
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1253
apply(nominal_induct R avoiding: c y x z M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1254
apply(auto split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1255
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1256
apply(auto simp add: alpha fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1257
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1258
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1259
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1260
apply(rule_tac x="[(name,z)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1261
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1262
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1263
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1264
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1265
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1266
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1267
apply(auto simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1268
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1269
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1270
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1271
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1272
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1273
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1274
lemma nrename_NotR:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1275
  assumes a: "R[x\<turnstile>n>y] = NotR (z).N c" "z\<sharp>(R,x,y)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1276
  shows "\<exists>N'. (R = NotR (z).N' c) \<and> N'[x\<turnstile>n>y] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1277
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1278
apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1279
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1280
apply(rule_tac x="[(name,z)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1281
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1282
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1283
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1284
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1285
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1286
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1287
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1288
lemma nrename_NotL:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1289
  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N z" "c\<sharp>R" "z\<sharp>(x,y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1290
  shows "\<exists>N'. (R = NotL <c>.N' z) \<and> N'[x\<turnstile>n>y] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1291
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1292
apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1293
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1294
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1295
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1296
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1297
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1298
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1299
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1300
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1301
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1302
lemma nrename_NotL':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1303
  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" "c\<sharp>R" "x\<noteq>y" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1305
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1306
apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1307
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1308
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1309
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1310
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1311
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1312
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1313
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1314
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1315
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1316
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1317
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1318
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1319
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1320
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1321
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1322
lemma nrename_NotL_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1323
  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1324
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1325
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1326
apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1327
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1328
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1329
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1330
lemma nrename_AndL1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1331
  assumes a: "R[x\<turnstile>n>y] = AndL1 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1332
  shows "\<exists>N'. (R = AndL1 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1333
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1334
apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1335
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1336
apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1337
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1338
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1339
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1340
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1341
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1342
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1343
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1344
lemma nrename_AndL1':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1345
  assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1347
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1348
apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1349
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1350
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1351
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1352
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1353
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1354
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1355
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1356
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1357
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1358
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1359
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1360
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1361
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1362
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1363
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1364
lemma nrename_AndL1_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1365
  assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1366
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1367
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1368
apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1369
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1370
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1371
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1372
lemma nrename_AndL2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1373
  assumes a: "R[x\<turnstile>n>y] = AndL2 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1374
  shows "\<exists>N'. (R = AndL2 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1375
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1376
apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1377
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1378
apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1379
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1380
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1381
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1382
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1383
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1384
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1385
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1386
lemma nrename_AndL2':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1387
  assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1389
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1390
apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1391
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1392
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1393
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1394
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1395
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1396
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1397
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1398
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1399
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1400
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1401
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1402
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1403
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1404
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1405
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1406
lemma nrename_AndL2_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1407
  assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1408
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1409
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1410
apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1411
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1412
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1413
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1414
lemma nrename_AndR:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1417
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1418
apply(nominal_induct R avoiding: x y c d e M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1419
apply(auto split: if_splits simp add: trm.inject alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1420
apply(simp add: fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1421
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1422
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1423
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1424
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1425
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1426
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1427
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1428
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1429
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1430
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1431
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1432
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1433
apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1434
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1435
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1436
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1437
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1438
lemma nrename_OrR1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1439
  assumes a: "R[x\<turnstile>n>y] = OrR1 <c>.N d" "c\<sharp>(R,d)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1440
  shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1441
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1442
apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1443
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1444
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1445
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1446
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1447
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1448
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1449
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1450
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1451
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1452
lemma nrename_OrR2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1453
  assumes a: "R[x\<turnstile>n>y] = OrR2 <c>.N d" "c\<sharp>(R,d)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1454
  shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1455
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1456
apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1457
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1458
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1459
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1460
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1461
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1462
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1463
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1464
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1465
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1466
lemma nrename_OrL:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1469
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1470
apply(nominal_induct R avoiding: u v x y z M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1471
apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1472
apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1473
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1474
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1475
apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1476
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1477
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1478
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1479
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1480
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1481
apply(drule_tac s="trm2[u\<turnstile>n>v]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1482
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1483
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1484
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1485
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1486
lemma nrename_OrL':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1490
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1491
apply(nominal_induct R avoiding: y x u v w M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1492
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1493
apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1494
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1495
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1496
apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1497
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1498
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1499
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1500
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1501
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1502
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1503
apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1504
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1505
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1506
apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1507
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1508
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1509
apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1510
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1511
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1512
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1513
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1514
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1515
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1516
apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1517
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1518
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1519
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1520
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1521
lemma nrename_OrL_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1522
  assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1523
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1524
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1525
apply(nominal_induct R avoiding: y x w u v M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1526
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1527
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1528
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1529
lemma nrename_ImpL:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1532
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1533
apply(nominal_induct R avoiding: u x c y z M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1534
apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1535
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1536
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1537
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1538
apply(rule_tac x="[(name1,u)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1539
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1540
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1541
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1542
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1543
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1544
apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1545
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1546
apply(simp add: eqvts calc_atm fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1547
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1548
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1549
lemma nrename_ImpL':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1553
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1554
apply(nominal_induct R avoiding: y x u c w M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1555
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1556
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1557
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1558
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1559
apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1560
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1561
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1562
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1563
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1564
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1565
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1566
apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1567
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1568
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1569
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1570
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1571
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1572
apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1573
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1574
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1575
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1576
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1577
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1578
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1579
apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1580
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1581
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1582
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1583
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1584
lemma nrename_ImpL_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1585
  assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1586
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1587
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1588
apply(nominal_induct R avoiding: y x w u c M N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1589
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1590
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1591
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1592
lemma nrename_ImpR:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1593
  assumes a: "R[u\<turnstile>n>v] = ImpR (x).<c>.N d" "c\<sharp>(R,d)" "x\<sharp>(R,u,v)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1594
  shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[u\<turnstile>n>v] = N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1595
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1596
apply(nominal_induct R avoiding: u v x c d N rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1597
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1598
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1599
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1600
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1601
apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1602
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1603
apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1604
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1605
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1606
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1607
apply(simp add: eqvts calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1608
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1609
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1610
lemma nrename_credu:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  1611
  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>c M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  1612
  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>c M0"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1613
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1614
apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: c_redu.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1615
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1616
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1617
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1618
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1619
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1620
apply(rule_tac x="M'{a:=(x).N'}" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1621
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1622
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1623
apply(rule c_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1624
apply(auto dest: not_fic_nrename)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1625
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1626
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1627
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1628
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1629
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1630
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1631
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1632
apply(rule_tac x="N'{x:=<a>.M'}" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1633
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1634
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1635
apply(rule c_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1636
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1637
apply(drule_tac x="xa" and y="y" in fin_nrename)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1638
apply(auto simp add: fresh_prod abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1639
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1640
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1641
lemma nrename_ax2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1642
  assumes a: "N[x\<turnstile>n>y] = Ax z c"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1643
  shows "\<exists>z. N = Ax z c"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1644
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1645
apply(nominal_induct N avoiding: x y rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1646
apply(auto split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1647
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1648
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1649
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1650
lemma fic_nrename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1651
  assumes a: "fic (M[x\<turnstile>n>y]) c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1652
  shows "fic M c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1653
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1654
apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1655
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1656
           split: if_splits)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1657
apply(auto dest: nrename_fresh_interesting2 simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1658
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1659
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1660
lemma nrename_lredu:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  1661
  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>l M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  1662
  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>l M0"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1663
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1664
apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: l_redu.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1665
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1666
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1667
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1668
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1669
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1670
apply(case_tac "xa=y")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1671
apply(simp add: nrename_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1672
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1673
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1674
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1675
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1676
apply(frule nrename_ax2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1677
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1678
apply(case_tac "z=xa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1679
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1680
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1681
apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1682
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1683
apply(rule crename_interesting3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1684
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1685
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1686
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1687
apply(auto dest: fic_nrename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1688
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1689
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1690
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1691
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1692
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1693
apply(case_tac "xa=ya")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1694
apply(simp add: nrename_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1695
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1696
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1697
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1698
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1699
apply(frule nrename_ax2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1700
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1701
apply(case_tac "z=xa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1702
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1703
apply(rule_tac x="N'[x\<turnstile>n>xa]" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1704
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1705
apply(rule nrename_interesting1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1706
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1707
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1708
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1709
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1710
apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1711
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1712
apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1713
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1714
apply(rule nrename_interesting2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1715
apply(simp_all)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1716
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1717
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1718
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1719
apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1720
(* LNot *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1721
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1722
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1723
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1724
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1725
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1726
apply(drule nrename_NotR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1727
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1728
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1729
apply(drule nrename_NotL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1730
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1731
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1732
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1733
apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1734
apply(simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1735
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1736
apply(auto simp add: fresh_prod fresh_atm intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1737
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1738
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1739
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1740
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1741
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1742
(* LAnd1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1743
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1744
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1745
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1746
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1747
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1748
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1749
apply(drule nrename_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1750
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1751
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1752
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1753
apply(drule nrename_AndL1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1754
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1755
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1756
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1757
apply(rule_tac x="Cut <a1>.M'a (x).N'b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1758
apply(simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1759
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1760
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1761
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1762
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1763
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1764
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1765
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1766
(* LAnd2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1767
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1768
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1769
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1770
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1771
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1772
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1773
apply(drule nrename_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1774
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1775
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1776
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1777
apply(drule nrename_AndL2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1778
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1779
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1780
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1781
apply(rule_tac x="Cut <a2>.N'a (x).N'b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1782
apply(simp add: fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1783
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1784
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1785
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1786
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1787
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1788
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1789
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1790
(* LOr1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1791
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1792
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1793
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1794
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1795
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1796
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1797
apply(drule nrename_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1798
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1799
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1800
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1801
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1802
apply(drule nrename_OrR1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1803
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1804
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1805
apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1806
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1807
apply(simp add: abs_fresh fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1808
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1809
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1810
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1811
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1812
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1813
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1814
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1815
(* LOr2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1816
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1817
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1818
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1819
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1820
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1821
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1822
apply(drule nrename_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1823
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1824
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1825
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1826
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1827
apply(drule nrename_OrR2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1828
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1829
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1830
apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1831
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1832
apply(simp add: abs_fresh fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1833
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1834
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1835
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1836
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1837
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1838
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1839
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1840
(* ImpL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1841
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1842
apply(drule sym) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1843
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1844
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1845
apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1846
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1847
apply(drule nrename_ImpL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1848
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1849
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1850
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1851
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1852
apply(drule nrename_ImpR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1853
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1854
apply(simp add: fresh_prod abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1855
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1856
apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1857
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1858
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1859
apply(rule l_redu.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1860
apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1861
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1862
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1863
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1864
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1865
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1866
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1867
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1868
lemma nrename_aredu:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  1869
  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>a M'" "x\<noteq>y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  1870
  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>a M0"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1871
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1872
apply(nominal_induct "M[x\<turnstile>n>y]" M' avoiding: M x y rule: a_redu.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1873
apply(drule  nrename_lredu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1874
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1875
apply(drule  nrename_credu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1876
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1877
(* Cut *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1878
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1879
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1880
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1881
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1882
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1883
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1884
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1885
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1886
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1887
apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1888
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1889
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1890
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1891
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1892
apply(drule nrename_fresh_interesting2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1893
apply(simp add: fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1894
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1895
apply(drule nrename_fresh_interesting1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1896
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1897
apply(simp add: fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1898
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1899
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1900
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1901
apply(drule nrename_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1902
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1903
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1904
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1905
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1906
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1907
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1908
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1909
apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1910
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1911
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1912
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1913
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1914
apply(simp add: fresh_a_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1915
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1916
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1917
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1918
(* NotL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1919
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1920
apply(frule nrename_NotL_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1921
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1922
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1923
apply(drule nrename_NotL')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1924
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1925
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1926
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1927
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1928
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1929
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1930
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1931
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1932
apply(rule_tac x="NotL <a>.M0 x" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1933
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1934
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1935
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1936
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1937
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1938
apply(drule_tac x="x" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1939
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1940
apply(rule_tac x="NotL <a>.M0 xa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1941
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1942
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1943
(* NotR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1944
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1945
apply(drule nrename_NotR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1946
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1947
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1948
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1949
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1950
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1951
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1952
apply(rule_tac x="NotR (x).M0 a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1953
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1954
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1955
(* AndR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1956
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1957
apply(drule nrename_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1958
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1959
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1960
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1961
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1962
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1963
apply(drule_tac x="x" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1964
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1965
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1966
apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1967
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1968
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1969
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1970
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1971
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1972
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1973
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1974
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1975
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1976
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1977
apply(drule nrename_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1978
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1979
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1980
apply(auto simp add: fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1981
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1982
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1983
apply(drule_tac x="x" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1984
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1985
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1986
apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1987
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1988
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1989
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1990
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1991
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1992
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1993
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1994
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1995
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1996
(* AndL1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1997
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1998
apply(frule nrename_AndL1_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1999
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2000
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2001
apply(drule nrename_AndL1')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2002
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2003
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2004
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2005
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2006
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2007
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2008
apply(drule_tac x="ya" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2009
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2010
apply(rule_tac x="AndL1 (x).M0 y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2011
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2012
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2013
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2014
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2015
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2016
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2017
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2018
apply(rule_tac x="AndL1 (x).M0 xa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2019
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2020
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2021
(* AndL2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2022
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2023
apply(frule nrename_AndL2_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2024
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2025
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2026
apply(drule nrename_AndL2')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2027
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2028
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2029
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2030
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2031
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2032
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2033
apply(drule_tac x="ya" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2034
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2035
apply(rule_tac x="AndL2 (x).M0 y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2036
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2037
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2038
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2039
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2040
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2041
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2042
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2043
apply(rule_tac x="AndL2 (x).M0 xa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2044
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2045
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2046
(* OrL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2047
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2048
apply(frule nrename_OrL_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2049
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2050
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2051
apply(drule nrename_OrL')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2052
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2053
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2054
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2055
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2056
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2057
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2058
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2059
apply(drule_tac x="ya" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2060
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2061
apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2062
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2063
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2064
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2065
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2066
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2067
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2068
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2069
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2070
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2071
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2072
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2073
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2074
apply(drule_tac x="z" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2075
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2076
apply(rule_tac x="OrL (x).M0 (y).N' xa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2077
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2078
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2079
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2080
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2081
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2082
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2083
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2084
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2085
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2086
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2087
apply(frule nrename_OrL_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2088
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2089
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2090
apply(drule nrename_OrL')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2091
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2092
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2093
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2094
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2095
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2096
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2097
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2098
apply(drule_tac x="ya" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2099
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2100
apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2101
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2102
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2103
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2104
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2105
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2106
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2107
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2108
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2109
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2110
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2111
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2112
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2113
apply(drule_tac x="z" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2114
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2115
apply(rule_tac x="OrL (x).M' (y).M0 xa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2116
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2117
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2118
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2119
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2120
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2121
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2122
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2123
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2124
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2125
(* OrR1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2126
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2127
apply(drule nrename_OrR1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2128
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2129
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2130
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2131
apply(drule_tac x="x" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2132
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2133
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2134
apply(rule_tac x="OrR1 <a>.M0 b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2135
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2136
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2137
(* OrR2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2138
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2139
apply(drule nrename_OrR2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2140
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2141
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2142
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2143
apply(drule_tac x="x" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2144
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2145
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2146
apply(rule_tac x="OrR2 <a>.M0 b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2147
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2148
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2149
(* ImpL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2150
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2151
apply(frule nrename_ImpL_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2152
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2153
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2154
apply(drule nrename_ImpL')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2155
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2156
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2157
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2158
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2159
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2160
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2161
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2162
apply(drule_tac x="ya" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2163
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2164
apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2165
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2166
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2167
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2168
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2169
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2170
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2171
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2172
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2173
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2174
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2175
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2176
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2177
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2178
apply(rule_tac x="ImpL <a>.M0 (x).N' xa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2179
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2180
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2181
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2182
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2183
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2184
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2185
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2186
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2187
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2188
apply(frule nrename_ImpL_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2189
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2190
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2191
apply(drule nrename_ImpL')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2192
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2193
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2194
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2195
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2196
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2197
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2198
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2199
apply(drule_tac x="ya" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2200
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2201
apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2202
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2203
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2204
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2205
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2206
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2207
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2208
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2209
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2210
apply(drule_tac x="N'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2211
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2212
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2213
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2214
apply(rule_tac x="ImpL <a>.M' (x).M0 xa" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2215
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2216
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2217
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2218
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2219
apply(auto intro: fresh_a_redu)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2220
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2221
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2222
(* ImpR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2223
apply(drule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2224
apply(drule nrename_ImpR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2225
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2226
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2227
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2228
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2229
apply(drule_tac x="xa" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2230
apply(drule_tac x="y" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2231
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2232
apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2233
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2234
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2235
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2236
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2237
lemma SNa_preserved_renaming2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2238
  assumes a: "SNa N"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2239
  shows "SNa (N[x\<turnstile>n>y])"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2240
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2241
apply(induct rule: SNa_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2242
apply(case_tac "x=y")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2243
apply(simp add: nrename_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2244
apply(rule SNaI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2245
apply(drule nrename_aredu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2246
apply(blast)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2247
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2248
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2249
text {* helper-stuff to set up the induction *}
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2250
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2251
abbreviation
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2252
  SNa_set :: "trm set"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2253
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2254
  "SNa_set \<equiv> {M. SNa M}"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2255
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2256
abbreviation
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2257
  A_Redu_set :: "(trm\<times>trm) set"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2258
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  2259
 "A_Redu_set \<equiv> {(N,M)| M N. M \<longrightarrow>\<^sub>a N}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2260
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2261
lemma SNa_elim:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2262
  assumes a: "SNa M"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  2263
  shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^sub>a N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2264
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2265
by (induct rule: SNa.induct) (blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2266
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2267
lemma wf_SNa_restricted:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2268
  shows "wf (A_Redu_set \<inter> (UNIV <*> SNa_set))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2269
apply(unfold wf_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2270
apply(intro strip)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2271
apply(case_tac "SNa x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2272
apply(simp (no_asm_use))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2273
apply(drule_tac P="P" in SNa_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2274
apply(erule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2275
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2276
(* other case *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2277
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2278
apply(erule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2279
apply(fast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2280
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2281
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2282
definition SNa_Redu :: "(trm \<times> trm) set" where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2283
  "SNa_Redu \<equiv> A_Redu_set \<inter> (UNIV <*> SNa_set)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2284
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2285
lemma wf_SNa_Redu:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2286
  shows "wf SNa_Redu"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2287
apply(unfold SNa_Redu_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2288
apply(rule wf_SNa_restricted)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2289
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2290
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2291
lemma wf_measure_triple:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2292
shows "wf ((measure size) <*lex*> SNa_Redu <*lex*> SNa_Redu)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2293
by (auto intro: wf_SNa_Redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2294
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2295
lemma my_wf_induct_triple: 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2296
 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2297
 and     b: "\<And>x. \<lbrakk>\<And>y. ((fst y,fst (snd y),snd (snd y)),(fst x,fst (snd x), snd (snd x))) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2298
                                    \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y\<rbrakk> \<Longrightarrow> P x"  
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2299
 shows "P x"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2300
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2301
apply(induct x rule: wf_induct_rule)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2302
apply(rule b)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2303
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2304
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2305
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2306
lemma my_wf_induct_triple': 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2307
 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2309
             \<Longrightarrow> P (x1,x2,x3)"  
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2310
 shows "P (x1,x2,x3)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2311
apply(rule_tac my_wf_induct_triple[OF a])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2312
apply(case_tac x rule: prod.exhaust)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2313
apply(simp)
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55413
diff changeset
  2314
apply(rename_tac p a b)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2315
apply(case_tac b)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2316
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2317
apply(rule b)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2318
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2319
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2320
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2321
lemma my_wf_induct_triple'': 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2322
 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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>
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2324
               \<Longrightarrow> P x1 x2 x3"  
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2325
 shows "P x1 x2 x3"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2326
apply(rule_tac my_wf_induct_triple'[where P="\<lambda>(x1,x2,x3). P x1 x2 x3", simplified])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2327
apply(rule a)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2328
apply(rule b)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2329
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2330
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2331
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2332
lemma excluded_m:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2333
  assumes a: "<a>:M \<in> (\<parallel><B>\<parallel>)" "(x):N \<in> (\<parallel>(B)\<parallel>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2334
  shows "(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2335
      \<or>\<not>(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2336
by (blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2337
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2338
lemma tricky_subst:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2339
  assumes a1: "b\<sharp>(c,N)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2340
  and     a2: "z\<sharp>(x,P)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2341
  and     a3: "M\<noteq>Ax z b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2342
  shows "(Cut <c>.N (z).M){b:=(x).P} = Cut <c>.N (z).(M{b:=(x).P})"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2343
using a1 a2 a3
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2344
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2345
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2346
apply(subgoal_tac "Cut <c>.N (z).M = Cut <ca>.([(ca,c)]\<bullet>N) (z).M")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2347
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2348
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2349
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2350
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2351
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2352
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2353
apply(subgoal_tac "b\<sharp>([(ca,c)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2354
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2355
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2356
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2357
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2358
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2359
apply(simp add: alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2360
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2361
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2362
text {* 3rd lemma *}
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2363
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2364
lemma CUT_SNa_aux:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2365
  assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2366
  and     a2: "SNa M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2367
  and     a3: "(x):N \<in> (\<parallel>(B)\<parallel>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2368
  and     a4: "SNa N"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2369
  shows   "SNa (Cut <a>.M (x).N)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2370
using a1 a2 a3 a4
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2371
apply(induct B M N arbitrary: a x rule: my_wf_induct_triple''[OF wf_measure_triple])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2372
apply(rule SNaI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2373
apply(drule Cut_a_redu_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2374
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2375
(* left-inner reduction *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2376
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2377
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2378
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2379
apply(drule_tac x="x1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2380
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2381
apply(drule_tac x="x3" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2382
apply(drule conjunct2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2383
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2384
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2385
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2386
apply(rule disjI1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2387
apply(simp add: SNa_Redu_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2388
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2389
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2390
apply(simp add: CANDs_preserved_single)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2391
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2392
apply(simp add: a_preserves_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2393
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2394
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2395
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2396
(* right-inner reduction *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2397
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2398
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2399
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2400
apply(drule_tac x="x1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2401
apply(drule_tac x="x2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2402
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2403
apply(drule conjunct2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2404
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2405
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2406
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2407
apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2408
apply(simp add: SNa_Redu_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2409
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2410
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2411
apply(simp add: CANDs_preserved_single)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2412
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2413
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2414
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2415
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2416
apply(simp add: CANDs_preserved_single)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2417
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2418
apply(simp add: a_preserves_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2419
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2420
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2421
(******** c-reduction *********)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2422
apply(drule Cut_c_redu_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2423
(* c-left reduction*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2424
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2425
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2426
apply(frule_tac B="x1" in fic_CANDS)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2427
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2428
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2429
(* in AXIOMSc *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2430
apply(simp add: AXIOMSc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2431
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2432
apply(simp add: ctrm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2433
apply(simp add: alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2434
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2435
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2436
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2437
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2438
apply(subgoal_tac "fic (Ax y b) b")(*A*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2439
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2440
(*A*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2441
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2442
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2443
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2444
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2445
apply(subgoal_tac "fic (Ax ([(a,aa)]\<bullet>y) a) a")(*B*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2446
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2447
(*B*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2448
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2449
(* in BINDINGc *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2450
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2451
apply(drule BINDINGc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2452
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2453
(* c-right reduction*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2454
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2455
apply(frule_tac B="x1" in fin_CANDS)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2456
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2457
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2458
(* in AXIOMSc *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2459
apply(simp add: AXIOMSn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2460
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2461
apply(simp add: ntrm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2462
apply(simp add: alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2463
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2464
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2465
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2466
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2467
apply(subgoal_tac "fin (Ax xa b) xa")(*A*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2468
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2469
(*A*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2470
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2471
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2472
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2473
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2474
apply(subgoal_tac "fin (Ax x ([(x,xa)]\<bullet>b)) x")(*B*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2475
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2476
(*B*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2477
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2478
(* in BINDINGc *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2479
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2480
apply(drule BINDINGn_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2481
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2482
(*********** l-reductions ************)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2483
apply(drule Cut_l_redu_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2484
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2485
(* ax1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2486
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2487
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2488
apply(simp add: SNa_preserved_renaming1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2489
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2490
(* ax2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2491
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2492
apply(simp add: SNa_preserved_renaming2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2493
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2494
(* LNot *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2495
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2496
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2497
apply(frule_tac excluded_m)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2498
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2499
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2500
(* one of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2501
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2502
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2503
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2504
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2505
apply(drule BINDINGc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2506
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2507
apply(drule_tac x="NotL <b>.N' x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2508
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2509
apply(simp add: better_NotR_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2510
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2511
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.NotR (y).M'a a' (x).NotL <b>.N' x) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2512
                   =  Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2513
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2515
apply(simp only: a_preserves_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2516
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2517
apply(rule better_LNot_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2518
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2519
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2520
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2521
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2522
(* other case of in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2523
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2524
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2525
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2526
apply(drule BINDINGn_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2527
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2528
apply(drule_tac x="NotR (y).M'a a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2529
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2530
apply(simp add: better_NotL_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2531
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2532
apply(subgoal_tac "fresh_fun (\<lambda>x'. Cut <a>.NotR (y).M'a a (x').NotL <b>.N' x') 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2533
                   = Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2534
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2536
apply(simp only: a_preserves_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2537
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2538
apply(rule better_LNot_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2539
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2540
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2541
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2542
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2543
(* none of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2544
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2545
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2546
apply(frule CAND_NotR_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2547
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2548
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2549
apply(frule CAND_NotL_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2550
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2551
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2552
apply(simp only: ty.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2553
apply(drule_tac x="B'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2554
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2555
apply(drule_tac x="M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2556
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2557
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2558
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2559
apply(drule_tac x="b" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2560
apply(rotate_tac 13)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2561
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2562
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2563
apply(rotate_tac 13)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2564
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2565
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2566
apply(drule_tac x="y" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2567
apply(rotate_tac 13)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2568
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2569
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2570
apply(rotate_tac 13)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2571
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2572
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2573
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2574
(* LAnd1 case *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2575
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2576
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2577
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2578
apply(frule_tac excluded_m)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2579
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2580
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2581
(* one of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2582
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2583
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2584
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2585
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2586
apply(drule BINDINGc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2587
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2588
apply(drule_tac x="AndL1 (y).N' x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2589
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2590
apply(simp add: better_AndR_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2591
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2592
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL1 (y).N' x) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2593
                   = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2594
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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'")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2596
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2597
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2598
apply(rule better_LAnd1_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2599
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2600
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2601
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2602
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2603
(* other case of in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2604
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2605
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2606
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2607
apply(drule BINDINGn_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2608
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2609
apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2610
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2611
apply(simp add: better_AndL1_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2612
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2613
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL1 (y).N' z') 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2614
                   = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2615
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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'")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2617
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2618
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2619
apply(rule better_LAnd1_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2620
apply(simp add: abs_fresh fresh_prod fresh_atm) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2621
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2622
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2623
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2624
(* none of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2625
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2626
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2627
apply(frule CAND_AndR_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2628
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2629
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2630
apply(frule CAND_AndL1_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2631
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2632
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2633
apply(simp only: ty.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2634
apply(drule_tac x="B1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2635
apply(drule_tac x="M1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2636
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2637
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2638
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2639
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2640
apply(drule_tac x="b" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2641
apply(rotate_tac 14)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2642
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2643
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2644
apply(rotate_tac 14)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2645
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2646
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2647
apply(drule_tac x="y" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2648
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2649
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2650
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2651
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2652
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2653
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2654
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2655
(* LAnd2 case *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2656
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2657
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2658
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2659
apply(frule_tac excluded_m)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2660
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2661
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2662
(* one of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2663
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2664
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2665
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2666
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2667
apply(drule BINDINGc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2668
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2669
apply(drule_tac x="AndL2 (y).N' x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2670
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2671
apply(simp add: better_AndR_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2672
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2673
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL2 (y).N' x) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2674
                   = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2675
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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'")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2677
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2678
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2679
apply(rule better_LAnd2_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2680
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2681
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2682
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2683
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2684
(* other case of in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2685
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2686
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2687
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2688
apply(drule BINDINGn_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2689
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2690
apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2691
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2692
apply(simp add: better_AndL2_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2693
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2694
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL2 (y).N' z') 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2695
                   = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2696
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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'")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2698
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2699
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2700
apply(rule better_LAnd2_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2701
apply(simp add: abs_fresh fresh_prod fresh_atm) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2702
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2703
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2704
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2705
(* none of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2706
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2707
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2708
apply(frule CAND_AndR_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2709
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2710
apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2711
apply(frule CAND_AndL2_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2712
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2713
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2714
apply(simp only: ty.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2715
apply(drule_tac x="B2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2716
apply(drule_tac x="M2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2717
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2718
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2719
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2720
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2721
apply(drule_tac x="c" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2722
apply(rotate_tac 14)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2723
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2724
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2725
apply(rotate_tac 14)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2726
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2727
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2728
apply(drule_tac x="y" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2729
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2730
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2731
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2732
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2733
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2734
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2735
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2736
(* LOr1 case *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2737
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2738
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2739
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2740
apply(frule_tac excluded_m)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2741
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2742
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2743
(* one of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2744
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2745
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2746
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2747
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2748
apply(drule BINDINGc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2749
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2750
apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2751
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2752
apply(simp add: better_OrR1_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2753
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2754
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2755
                   = Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2756
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2758
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2759
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2760
apply(rule better_LOr1_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2761
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2762
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2763
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2764
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2765
(* other case of in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2766
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2767
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2768
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2769
apply(drule BINDINGn_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2770
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2771
apply(drule_tac x="OrR1 <b>.N' a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2772
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2773
apply(simp add: better_OrL_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2774
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2775
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR1 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2776
                   = Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2777
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2779
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2780
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2781
apply(rule better_LOr1_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2782
apply(simp add: abs_fresh fresh_prod fresh_atm) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2783
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2784
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2785
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2786
(* none of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2787
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2788
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2789
apply(frule CAND_OrR1_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2790
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2791
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2792
apply(frule CAND_OrL_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2793
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2794
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2795
apply(simp only: ty.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2796
apply(drule_tac x="B1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2797
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2798
apply(drule_tac x="M1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2799
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2800
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2801
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2802
apply(drule_tac x="b" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2803
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2804
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2805
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2806
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2807
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2808
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2809
apply(drule_tac x="z" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2810
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2811
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2812
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2813
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2814
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2815
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2816
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2817
(* LOr2 case *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2818
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2819
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2820
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2821
apply(frule_tac excluded_m)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2822
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2823
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2824
(* one of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2825
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2826
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2827
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2828
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2829
apply(drule BINDINGc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2830
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2831
apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2832
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2833
apply(simp add: better_OrR2_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2834
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2835
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2836
                   = Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2837
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2839
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2840
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2841
apply(rule better_LOr2_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2842
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2843
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2844
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2845
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2846
(* other case of in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2847
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2848
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2849
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2850
apply(drule BINDINGn_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2851
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2852
apply(drule_tac x="OrR2 <b>.N' a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2853
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2854
apply(simp add: better_OrL_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2855
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2856
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR2 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2857
                   = Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2858
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  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")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2860
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2861
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2862
apply(rule better_LOr2_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2863
apply(simp add: abs_fresh fresh_prod fresh_atm) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2864
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2865
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2866
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2867
(* none of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2868
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2869
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2870
apply(frule CAND_OrR2_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2871
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2872
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2873
apply(frule CAND_OrL_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2874
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2875
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2876
apply(simp only: ty.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2877
apply(drule_tac x="B2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2878
apply(drule_tac x="N'" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2879
apply(drule_tac x="M2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2880
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2881
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2882
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2883
apply(drule_tac x="b" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2884
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2885
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2886
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2887
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2888
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2889
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2890
apply(drule_tac x="y" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2891
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2892
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2893
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2894
apply(rotate_tac 15)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2895
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2896
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2897
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2898
(* LImp case *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2899
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2900
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2901
apply(frule_tac excluded_m)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2902
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2903
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2904
(* one of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2905
apply(erule disjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2906
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2907
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2908
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2909
apply(drule BINDINGc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2910
apply(drule_tac x="x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2911
apply(drule_tac x="ImpL <c>.N1 (y).N2 x" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2912
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2913
apply(simp add: better_ImpR_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2914
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2915
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.ImpR (z).<b>.M'a a' (x).ImpL <c>.N1 (y).N2 x)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2916
                   = Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2917
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  2918
apply(subgoal_tac "Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x \<longrightarrow>\<^sub>a 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2919
                                                          Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2920
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2921
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2922
apply(rule better_LImp_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2923
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2924
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2925
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2926
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2927
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2928
(* other case of in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2929
apply(drule fin_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2930
apply(drule fic_elims)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2931
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2932
apply(drule BINDINGn_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2933
apply(drule_tac x="a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2934
apply(drule_tac x="ImpR (z).<b>.M'a a" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2935
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2936
apply(simp add: better_ImpL_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2937
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2938
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.ImpR (z).<b>.M'a a (z').ImpL <c>.N1 (y).N2 z')
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2939
                   = Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2940
apply(simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  2941
apply(subgoal_tac "Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca \<longrightarrow>\<^sub>a 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2942
                                                          Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2943
apply(auto intro: a_preserves_SNa)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2944
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2945
apply(rule better_LImp_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2946
apply(simp add: abs_fresh fresh_prod fresh_atm) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2947
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2948
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2949
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2950
apply(simp add: abs_fresh abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2951
apply(simp add: abs_fresh abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2952
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2953
(* none of them in BINDING *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2954
apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2955
apply(frule CAND_ImpL_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2956
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2957
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2958
apply(frule CAND_ImpR_elim) (* check here *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2959
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2960
apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2961
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2962
apply(simp only: ty.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2963
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2964
apply(case_tac "M'a=Ax z b")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2965
(* case Ma = Ax z b *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2967
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2968
apply(drule_tac x="c" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2969
apply(drule_tac x="N1" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2970
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2971
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2972
apply(drule_tac x="B2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2973
apply(drule_tac x="M'a{z:=<c>.N1}" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2974
apply(drule_tac x="N2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2975
apply(drule conjunct1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2976
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2977
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2978
apply(rotate_tac 17)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2979
apply(drule_tac x="b" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2980
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2981
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2982
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2983
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2984
apply(rotate_tac 17)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2985
apply(drule_tac x="y" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2986
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2987
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2988
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2989
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2990
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2991
(* case Ma \<noteq> Ax z b *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2992
apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> \<parallel><B2>\<parallel>") (* lemma *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2993
apply(frule_tac meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2994
apply(drule_tac x="B2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2995
apply(drule_tac x="Cut <c>.N1 (z).M'a" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2996
apply(drule_tac x="N2" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2997
apply(erule conjE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2998
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2999
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3000
apply(rotate_tac 20)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3001
apply(drule_tac x="b" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3002
apply(rotate_tac 20)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3003
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3004
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3005
apply(rotate_tac 20)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3006
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3007
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3008
apply(rotate_tac 20)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3009
apply(drule_tac x="y" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3010
apply(rotate_tac 20)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3011
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3012
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3013
apply(rotate_tac 20)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3014
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3015
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3016
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3017
(* lemma *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3018
apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> BINDINGc B2 (\<parallel>(B2)\<parallel>)") (* second lemma *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3019
apply(simp add: BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3020
(* second lemma *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3021
apply(simp (no_asm) add: BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3022
apply(rule exI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3023
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3024
apply(rule refl)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3025
apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3026
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3027
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3028
apply(rule_tac t="Cut <c>.N1 (z).M'a" and s="Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)" in subst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3029
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3030
apply(rule_tac t="(Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)){b:=(xa).P}" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3031
           and s="Cut <c>.N1 (ca).(([(ca,z)]\<bullet>M'a){b:=(xa).P})" in subst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3032
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3033
apply(rule tricky_subst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3034
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3035
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3036
apply(clarify)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3037
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3038
apply(simp add: calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3039
apply(drule_tac x="B1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3040
apply(drule_tac x="N1" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3041
apply(drule_tac x="([(ca,z)]\<bullet>M'a){b:=(xa).P}" in meta_spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3042
apply(drule conjunct1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3043
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3044
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3045
apply(rotate_tac 19)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3046
apply(drule_tac x="c" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3047
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3048
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3049
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3050
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3051
apply(rotate_tac 19)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3052
apply(drule_tac x="ca" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3053
apply(subgoal_tac "(ca):([(ca,z)]\<bullet>M'a){b:=(xa).P} \<in> \<parallel>(B1)\<parallel>")(*A*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3054
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3055
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3056
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3057
apply(simp add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3058
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3059
(*A*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3060
apply(drule_tac x="[(ca,z)]\<bullet>xa" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3061
apply(drule_tac x="[(ca,z)]\<bullet>P" in spec)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3062
apply(rotate_tac 19)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3063
apply(simp add: fresh_prod fresh_left)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3064
apply(drule mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3065
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3066
apply(auto simp add: calc_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3067
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3068
apply(auto simp add: calc_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3069
apply(drule_tac pi="[(ca,z)]" and x="(xa):P" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3070
apply(simp add: CAND_eqvt_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3071
apply(drule_tac pi="[(ca,z)]" and X="\<parallel>(B1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3072
apply(simp add: CAND_eqvt_name csubst_eqvt)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3073
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3074
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3075
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3076
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3077
(* parallel substitution *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3078
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3079
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3080
lemma CUT_SNa:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3081
  assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3082
  and     a2: "(x):N \<in> (\<parallel>(B)\<parallel>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3083
  shows   "SNa (Cut <a>.M (x).N)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3084
using a1 a2
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3085
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3086
apply(rule CUT_SNa_aux[OF a1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3087
apply(simp_all add: CANDs_imply_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3088
done 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3089
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3090
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3091
fun 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3092
 findn :: "(name\<times>coname\<times>trm) list\<Rightarrow>name\<Rightarrow>(coname\<times>trm) option"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3093
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3094
  "findn [] x = None"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3095
| "findn ((y,c,P)#\<theta>_n) x = (if y=x then Some (c,P) else findn \<theta>_n x)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3096
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3097
lemma findn_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3098
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3099
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3100
  shows "(pi1\<bullet>findn \<theta>_n x) = findn (pi1\<bullet>\<theta>_n) (pi1\<bullet>x)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3101
  and   "(pi2\<bullet>findn \<theta>_n x) = findn (pi2\<bullet>\<theta>_n) (pi2\<bullet>x)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3102
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3103
apply(auto simp add: perm_bij) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3104
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3105
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3106
lemma findn_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3107
  assumes a: "x\<sharp>\<theta>_n"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3108
  shows "findn \<theta>_n x = None"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3109
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3110
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3111
apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3112
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3113
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3114
fun 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3115
 findc :: "(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>(name\<times>trm) option"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3116
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3117
  "findc [] x = None"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3118
| "findc ((c,y,P)#\<theta>_c) a = (if a=c then Some (y,P) else findc \<theta>_c a)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3119
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3120
lemma findc_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3121
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3122
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3123
  shows "(pi1\<bullet>findc \<theta>_c a) = findc (pi1\<bullet>\<theta>_c) (pi1\<bullet>a)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3124
  and   "(pi2\<bullet>findc \<theta>_c a) = findc (pi2\<bullet>\<theta>_c) (pi2\<bullet>a)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3125
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3126
apply(auto simp add: perm_bij) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3127
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3128
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3129
lemma findc_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3130
  assumes a: "a\<sharp>\<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3131
  shows "findc \<theta>_c a = None"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3132
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3133
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3134
apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3135
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3136
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3137
abbreviation 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3138
 nmaps :: "(name\<times>coname\<times>trm) list \<Rightarrow> name \<Rightarrow> (coname\<times>trm) option \<Rightarrow> bool" ("_ nmaps _ to _" [55,55,55] 55) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3139
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3140
 "\<theta>_n nmaps x to P \<equiv> (findn \<theta>_n x) = P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3141
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3142
abbreviation 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3143
 cmaps :: "(coname\<times>name\<times>trm) list \<Rightarrow> coname \<Rightarrow> (name\<times>trm) option \<Rightarrow> bool" ("_ cmaps _ to _" [55,55,55] 55) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3144
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3145
 "\<theta>_c cmaps a to P \<equiv> (findc \<theta>_c a) = P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3146
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3147
lemma nmaps_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3148
  shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> a\<sharp>\<theta>_n \<Longrightarrow> a\<sharp>(c,P)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3149
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3150
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3151
apply(case_tac "aa=x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3152
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3153
apply(case_tac "aa=x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3154
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3155
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3156
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3157
lemma cmaps_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3158
  shows "\<theta>_c cmaps a to Some (y,P) \<Longrightarrow> x\<sharp>\<theta>_c \<Longrightarrow> x\<sharp>(y,P)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3159
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3160
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3161
apply(case_tac "a=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3162
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3163
apply(case_tac "a=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3164
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3165
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3166
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3167
lemma nmaps_false:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3168
  shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> x\<sharp>\<theta>_n \<Longrightarrow> False"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3169
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3170
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3171
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3172
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3173
lemma cmaps_false:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3174
  shows "\<theta>_c cmaps c to Some (x,P) \<Longrightarrow> c\<sharp>\<theta>_c \<Longrightarrow> False"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3175
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3176
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3177
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3178
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3179
fun 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3180
 lookupa :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3181
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3182
  "lookupa x a [] = Ax x a"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3184
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3185
lemma lookupa_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3186
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3187
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3188
  shows "(pi1\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_c)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3189
  and   "(pi2\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_c)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3190
apply -
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3191
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3192
apply(auto simp add: eqvts)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3193
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3194
apply(auto simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3195
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3196
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3197
lemma lookupa_fire:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3198
  assumes a: "\<theta>_c cmaps a to Some (y,P)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3199
  shows "(lookupa x a \<theta>_c) = Cut <a>.Ax x a (y).P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3200
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3201
apply(induct \<theta>_c arbitrary: x a y P)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3202
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3203
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3204
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3205
fun 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3206
 lookupb :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>trm\<Rightarrow>trm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3207
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3208
  "lookupb x a [] c P = Cut <c>.P (x).Ax x a"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3210
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3211
lemma lookupb_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3212
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3213
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3216
apply -
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3217
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3218
apply(auto simp add: eqvts)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3219
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3220
apply(auto simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3221
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3222
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3223
fun 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3224
  lookup :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3225
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3226
  "lookup x a [] \<theta>_c = lookupa x a \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3228
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3229
lemma lookup_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3230
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3231
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3234
apply -
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3235
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3236
apply(auto simp add: eqvts)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3237
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3238
apply(auto simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3239
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3240
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3241
fun 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3242
  lookupc :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3243
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3244
  "lookupc x a [] = Ax x a"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3245
| "lookupc x a ((y,c,P)#\<theta>_n) = (if x=y then P[c\<turnstile>c>a] else lookupc x a \<theta>_n)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3246
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3247
lemma lookupc_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3248
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3249
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3250
  shows "(pi1\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3251
  and   "(pi2\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3252
apply -
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3253
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3254
apply(auto simp add: eqvts)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3255
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3256
apply(auto simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3257
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3258
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3259
fun 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3260
  lookupd :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3261
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3262
  "lookupd x a [] = Ax x a"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3263
| "lookupd x a ((c,y,P)#\<theta>_c) = (if a=c then P[y\<turnstile>n>x] else lookupd x a \<theta>_c)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3264
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3265
lemma lookupd_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3266
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3267
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3268
  shows "(pi1\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3269
  and   "(pi2\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3270
apply -
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3271
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3272
apply(auto simp add: eqvts)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3273
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3274
apply(auto simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3275
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3276
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3277
lemma lookupa_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3278
  assumes a: "a\<sharp>\<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3279
  shows "lookupa y a \<theta>_c = Ax y a"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3280
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3281
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3282
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3283
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3284
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3285
lemma lookupa_csubst:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3286
  assumes a: "a\<sharp>\<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3287
  shows "Cut <a>.Ax y a (x).P = (lookupa y a \<theta>_c){a:=(x).P}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3288
using a by (simp add: lookupa_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3289
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3290
lemma lookupa_freshness:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3291
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3292
  and   x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3293
  shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupa y c \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3294
  and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupa y c \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3295
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3296
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3297
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3298
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3299
lemma lookupa_unicity:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3300
  assumes a: "lookupa x a \<theta>_c= Ax y b" "b\<sharp>\<theta>_c" "y\<sharp>\<theta>_c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3301
  shows "x=y \<and> a=b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3302
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3303
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3304
apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3305
apply(case_tac "a=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3306
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3307
apply(case_tac "a=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3308
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3309
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3310
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3311
lemma lookupb_csubst:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3312
  assumes a: "a\<sharp>(\<theta>_c,c,N)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3313
  shows "Cut <c>.N (x).P = (lookupb y a \<theta>_c c N){a:=(x).P}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3314
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3315
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3316
apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3317
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3318
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3319
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3320
apply(subgoal_tac "Cut <c>.N (y).Ax y a = Cut <caa>.([(caa,c)]\<bullet>N) (ca).Ax ca a")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3321
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3322
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3323
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3324
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3325
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3326
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3327
apply(subgoal_tac "a\<sharp>([(caa,c)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3328
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3329
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3330
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3331
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3332
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3333
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3334
apply(simp add: alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3335
apply(simp add: alpha calc_atm fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3336
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3337
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3338
lemma lookupb_freshness:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3339
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3340
  and   x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3341
  shows "a\<sharp>(\<theta>_c,c,b,P) \<Longrightarrow> a\<sharp>lookupb y c \<theta>_c b P"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3342
  and   "x\<sharp>(\<theta>_c,y,P) \<Longrightarrow> x\<sharp>lookupb y c \<theta>_c b P"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3343
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3344
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3345
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3346
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3347
lemma lookupb_unicity:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3348
  assumes a: "lookupb x a \<theta>_c c P = Ax y b" "b\<sharp>(\<theta>_c,c,P)" "y\<sharp>\<theta>_c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3349
  shows "x=y \<and> a=b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3350
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3351
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3352
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3353
apply(case_tac "a=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3354
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3355
apply(case_tac "a=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3356
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3357
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3358
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3359
lemma lookupb_lookupa:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3360
  assumes a: "x\<sharp>\<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3361
  shows "lookupb x c \<theta>_c a P = (lookupa x c \<theta>_c){x:=<a>.P}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3362
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3363
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3364
apply(auto simp add: fresh_list_cons fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3365
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3366
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3367
apply(subgoal_tac "Cut <c>.Ax x c (aa).b = Cut <ca>.Ax x ca (caa).([(caa,aa)]\<bullet>b)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3368
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3369
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3370
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3371
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3372
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3373
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3374
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3375
apply(subgoal_tac "x\<sharp>([(caa,aa)]\<bullet>b)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3376
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3377
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3378
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3379
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3380
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3381
apply(simp add: alpha calc_atm fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3382
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3383
apply(simp add: alpha calc_atm fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3384
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3385
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3386
lemma lookup_csubst:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3387
  assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3388
  shows "lookup y c \<theta>_n ((a,x,P)#\<theta>_c) = (lookup y c \<theta>_n \<theta>_c){a:=(x).P}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3389
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3390
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3391
apply(auto simp add: fresh_prod fresh_list_cons)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3392
apply(simp add: lookupa_csubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3393
apply(simp add: lookupa_freshness forget fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3394
apply(rule lookupb_csubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3395
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3396
apply(auto simp add: lookupb_freshness forget fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3397
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3398
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3399
lemma lookup_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3400
  assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3401
  shows "lookup x c \<theta>_n \<theta>_c = lookupa x c \<theta>_c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3402
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3403
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3404
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3405
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3406
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3407
lemma lookup_unicity:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3409
  shows "x=y \<and> a=b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3410
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3411
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3412
apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3413
apply(drule lookupa_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3414
apply(simp)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3415
apply(drule lookupa_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3416
apply(simp)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3417
apply(case_tac "x=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3418
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3419
apply(drule lookupb_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3420
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3421
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3422
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3423
apply(case_tac "x=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3424
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3425
apply(drule lookupb_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3426
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3427
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3428
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3429
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3430
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3431
lemma lookup_freshness:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3432
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3433
  and   x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3434
  shows "a\<sharp>(c,\<theta>_c,\<theta>_n) \<Longrightarrow> a\<sharp>lookup y c \<theta>_n \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3435
  and   "x\<sharp>(y,\<theta>_c,\<theta>_n) \<Longrightarrow> x\<sharp>lookup y c \<theta>_n \<theta>_c"   
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3436
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3437
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3438
apply(simp add: fresh_atm fresh_prod lookupa_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3439
apply(simp add: fresh_atm fresh_prod lookupa_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3440
apply(simp add: fresh_atm fresh_prod lookupb_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3441
apply(simp add: fresh_atm fresh_prod lookupb_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3442
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3443
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3444
lemma lookupc_freshness:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3445
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3446
  and   x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3447
  shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupc y c \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3448
  and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupc y c \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3449
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3450
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3451
apply(rule rename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3452
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3453
apply(rule rename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3454
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3455
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3456
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3457
lemma lookupc_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3458
  assumes a: "y\<sharp>\<theta>_n"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3459
  shows "lookupc y a \<theta>_n = Ax y a"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3460
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3461
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3462
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3463
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3464
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3465
lemma lookupc_nmaps:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3466
  assumes a: "\<theta>_n nmaps x to Some (c,P)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3467
  shows "lookupc x a \<theta>_n = P[c\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3468
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3469
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3470
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3471
done 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3472
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3473
lemma lookupc_unicity:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3474
  assumes a: "lookupc y a \<theta>_n = Ax x b" "x\<sharp>\<theta>_n"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3475
  shows "y=x"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3476
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3477
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3478
apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3479
apply(case_tac "y=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3480
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3481
apply(subgoal_tac "x\<sharp>(ba[aa\<turnstile>c>a])")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3482
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3483
apply(rule rename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3484
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3485
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3486
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3487
lemma lookupd_fresh:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3488
  assumes a: "a\<sharp>\<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3489
  shows "lookupd y a \<theta>_c = Ax y a"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3490
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3491
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3492
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3493
done 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3494
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3495
lemma lookupd_unicity:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3496
  assumes a: "lookupd y a \<theta>_c = Ax y b" "b\<sharp>\<theta>_c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3497
  shows "a=b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3498
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3499
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3500
apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3501
apply(case_tac "a=aa")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3502
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3503
apply(subgoal_tac "b\<sharp>(ba[aa\<turnstile>n>y])")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3504
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3505
apply(rule rename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3506
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3507
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3508
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3509
lemma lookupd_freshness:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3510
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3511
  and   x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3512
  shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupd y c \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3513
  and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupd y c \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3514
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3515
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3516
apply(rule rename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3517
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3518
apply(rule rename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3519
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3520
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3521
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3522
lemma lookupd_cmaps:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3523
  assumes a: "\<theta>_c cmaps a to Some (x,P)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3524
  shows "lookupd y a \<theta>_c = P[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3525
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3526
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3527
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3528
done 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3529
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3530
nominal_primrec (freshness_context: "\<theta>_n::(name\<times>coname\<times>trm)")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3531
  stn :: "trm\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3532
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3533
  "stn (Ax x a) \<theta>_n = lookupc x a \<theta>_n"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3535
| "x\<sharp>\<theta>_n \<Longrightarrow> stn (NotR (x).M a) \<theta>_n = (NotR (x).M a)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3536
| "a\<sharp>\<theta>_n \<Longrightarrow>stn (NotL <a>.M x) \<theta>_n = (NotL <a>.M x)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3538
| "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL1 (x).M z) \<theta>_n = (AndL1 (x).M z)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3539
| "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL2 (x).M z) \<theta>_n = (AndL2 (x).M z)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3540
| "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR1 <a>.M b) \<theta>_n = (OrR1 <a>.M b)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3541
| "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR2 <a>.M b) \<theta>_n = (OrR2 <a>.M b)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3545
apply(finite_guess)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3546
apply(rule TrueI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3547
apply(simp add: abs_fresh abs_supp fin_supp)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3548
apply(fresh_guess)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3549
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3550
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3551
nominal_primrec (freshness_context: "\<theta>_c::(coname\<times>name\<times>trm)")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3552
  stc :: "trm\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3553
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3554
  "stc (Ax x a) \<theta>_c = lookupd x a \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3556
| "x\<sharp>\<theta>_c \<Longrightarrow> stc (NotR (x).M a) \<theta>_c = (NotR (x).M a)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3557
| "a\<sharp>\<theta>_c \<Longrightarrow> stc (NotL <a>.M x) \<theta>_c = (NotL <a>.M x)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3559
| "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL1 (x).M z) \<theta>_c = (AndL1 (x).M z)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3560
| "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL2 (x).M z) \<theta>_c = (AndL2 (x).M z)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3561
| "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR1 <a>.M b) \<theta>_c = (OrR1 <a>.M b)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3562
| "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR2 <a>.M b) \<theta>_c = (OrR2 <a>.M b)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3566
apply(finite_guess)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3567
apply(rule TrueI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3568
apply(simp add: abs_fresh abs_supp fin_supp)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3569
apply(fresh_guess)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3570
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3571
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3572
lemma stn_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3573
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3574
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3575
  shows "(pi1\<bullet>(stn M \<theta>_n)) = stn (pi1\<bullet>M) (pi1\<bullet>\<theta>_n)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3576
  and   "(pi2\<bullet>(stn M \<theta>_n)) = stn (pi2\<bullet>M) (pi2\<bullet>\<theta>_n)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3577
apply -
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3578
apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3579
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3580
apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3581
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3582
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3583
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3584
lemma stc_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3585
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3586
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3587
  shows "(pi1\<bullet>(stc M \<theta>_c)) = stc (pi1\<bullet>M) (pi1\<bullet>\<theta>_c)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3588
  and   "(pi2\<bullet>(stc M \<theta>_c)) = stc (pi2\<bullet>M) (pi2\<bullet>\<theta>_c)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3589
apply -
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3590
apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3591
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3592
apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3593
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3594
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3595
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3596
lemma stn_fresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3597
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3598
  and   x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3599
  shows "a\<sharp>(\<theta>_n,M) \<Longrightarrow> a\<sharp>stn M \<theta>_n"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3600
  and   "x\<sharp>(\<theta>_n,M) \<Longrightarrow> x\<sharp>stn M \<theta>_n"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3601
apply(nominal_induct M avoiding: \<theta>_n a x rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3602
apply(auto simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3603
apply(rule lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3604
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3605
apply(rule lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3606
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3607
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3608
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3609
lemma stc_fresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3610
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3611
  and   x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3612
  shows "a\<sharp>(\<theta>_c,M) \<Longrightarrow> a\<sharp>stc M \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3613
  and   "x\<sharp>(\<theta>_c,M) \<Longrightarrow> x\<sharp>stc M \<theta>_c"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3614
apply(nominal_induct M avoiding: \<theta>_c a x rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3615
apply(auto simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3616
apply(rule lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3617
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3618
apply(rule lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3619
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3620
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3621
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 53015
diff changeset
  3622
lemma case_option_eqvt1[eqvt_force]:
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3623
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3624
  and   pi2::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3625
  and   B::"(name\<times>trm) option"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3626
  and   r::"trm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3627
  shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3628
              (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3629
  and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3630
              (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3631
apply(cases "B")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3632
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3633
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3634
apply(cases "B")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3635
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3636
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3637
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3638
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 53015
diff changeset
  3639
lemma case_option_eqvt2[eqvt_force]:
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3640
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3641
  and   pi2::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3642
  and   B::"(coname\<times>trm) option"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3643
  and   r::"trm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3644
  shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3645
              (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3646
  and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3647
              (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3648
apply(cases "B")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3649
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3650
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3651
apply(cases "B")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3652
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3653
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3654
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3655
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3656
nominal_primrec (freshness_context: "(\<theta>_n::(name\<times>coname\<times>trm) list,\<theta>_c::(coname\<times>name\<times>trm) list)")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3657
  psubst :: "(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm\<Rightarrow>trm" ("_,_<_>" [100,100,100] 100) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3658
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3659
  "\<theta>_n,\<theta>_c<Ax x a> = lookup x a \<theta>_n \<theta>_c" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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> = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3661
   Cut <a>.(if \<exists>x. M=Ax x a then stn M \<theta>_n else \<theta>_n,\<theta>_c<M>) 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3662
       (x).(if \<exists>a. N=Ax x a then stc N \<theta>_c else \<theta>_n,\<theta>_c<N>)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3663
| "x\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotR (x).M a> = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3664
  (case (findc \<theta>_c a) of 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3665
       Some (u,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).(\<theta>_n,\<theta>_c<M>) a' (u).P) 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3666
     | None \<Rightarrow> NotR (x).(\<theta>_n,\<theta>_c<M>) a)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3667
| "a\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotL <a>.M x> = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3668
  (case (findn \<theta>_n x) of 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3669
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>x'. Cut <c>.P (x').(NotL <a>.(\<theta>_n,\<theta>_c<M>) x')) 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3670
     | None \<Rightarrow> NotL <a>.(\<theta>_n,\<theta>_c<M>) x)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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>) = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3672
  (case (findc \<theta>_c c) of 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3674
     | None \<Rightarrow> AndR <a>.(\<theta>_n,\<theta>_c<M>) <b>.(\<theta>_n,\<theta>_c<N>) c)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3675
| "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL1 (x).M z>) = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3676
  (case (findn \<theta>_n z) of 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3677
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).(\<theta>_n,\<theta>_c<M>) z') 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3678
     | None \<Rightarrow> AndL1 (x).(\<theta>_n,\<theta>_c<M>) z)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3679
| "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL2 (x).M z>) = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3680
  (case (findn \<theta>_n z) of 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3681
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).(\<theta>_n,\<theta>_c<M>) z') 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3682
     | None \<Rightarrow> AndL2 (x).(\<theta>_n,\<theta>_c<M>) z)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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>) =
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3684
  (case (findn \<theta>_n z) of  
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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') 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3686
     | None \<Rightarrow> OrL (x).(\<theta>_n,\<theta>_c<M>) (u).(\<theta>_n,\<theta>_c<N>) z)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3687
| "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR1 <a>.M b>) = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3688
  (case (findc \<theta>_c b) of
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3689
       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3690
     | None \<Rightarrow> OrR1 <a>.(\<theta>_n,\<theta>_c<M>) b)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3691
| "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR2 <a>.M b>) = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3692
  (case (findc \<theta>_c b) of
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3693
       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3694
     | None \<Rightarrow> OrR2 <a>.(\<theta>_n,\<theta>_c<M>) b)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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>) = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3696
  (case (findc \<theta>_c b) of
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3697
       Some (z,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) a' (z).P)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3698
     | None \<Rightarrow> ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) b)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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>) = 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3700
  (case (findn \<theta>_n z) of
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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') 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3702
     | None \<Rightarrow> ImpL <a>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>) z)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3703
apply(finite_guess)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3704
apply(rule TrueI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3705
apply(simp add: abs_fresh stc_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3706
apply(simp add: abs_fresh stn_fresh)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3707
apply(case_tac "findc \<theta>_c x3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3708
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3709
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3710
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3711
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3712
apply(drule cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3713
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3714
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3715
apply(case_tac "findn \<theta>_n x3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3716
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3717
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3718
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3719
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3720
apply(drule nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3721
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3722
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3723
apply(case_tac "findc \<theta>_c x5")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3724
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3725
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3726
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3727
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3728
apply(drule cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3729
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3730
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3731
apply(case_tac "findc \<theta>_c x5")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3732
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3733
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3734
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3735
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3736
apply(drule_tac x="x3" in cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3737
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3738
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3739
apply(case_tac "findn \<theta>_n x3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3740
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3741
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3742
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3743
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3744
apply(drule nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3745
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3746
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3747
apply(case_tac "findn \<theta>_n x3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3748
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3749
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3750
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3751
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3752
apply(drule nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3753
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3754
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3755
apply(case_tac "findc \<theta>_c x3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3756
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3757
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3758
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3759
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3760
apply(drule cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3761
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3762
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3763
apply(case_tac "findc \<theta>_c x3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3764
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3765
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3766
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3767
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3768
apply(drule cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3769
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3770
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3771
apply(case_tac "findn \<theta>_n x5")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3772
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3773
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3774
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3775
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3776
apply(drule nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3777
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3778
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3779
apply(case_tac "findn \<theta>_n x5")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3780
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3781
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3782
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3783
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3784
apply(drule_tac a="x3" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3785
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3786
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3787
apply(case_tac "findc \<theta>_c x4")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3788
apply(simp add: abs_fresh abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3789
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3790
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3791
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3792
apply(drule cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3793
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3794
apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3795
apply(case_tac "findc \<theta>_c x4")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3796
apply(simp add: abs_fresh abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3797
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3798
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3799
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3800
apply(drule_tac x="x2" in cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3801
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3802
apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3803
apply(case_tac "findn \<theta>_n x5")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3804
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3805
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3806
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3807
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3808
apply(drule nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3809
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3810
apply(simp add: abs_fresh fresh_prod fresh_atm)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3811
apply(case_tac "findn \<theta>_n x5")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3812
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3813
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3814
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3815
apply(fresh_fun_simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3816
apply(drule_tac a="x3" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3817
apply(auto simp add: fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3818
apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3819
apply(fresh_guess)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3820
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3821
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3822
lemma case_cong:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3823
  assumes a: "B1=B2" "x1=x2" "y1=y2"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3825
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3826
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3827
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3828
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3829
lemma find_maps:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3830
  shows "\<theta>_c cmaps a to (findc \<theta>_c a)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3831
  and   "\<theta>_n nmaps x to (findn \<theta>_n x)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3832
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3833
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3834
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3835
lemma psubst_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3836
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3837
  and   pi2::"coname prm"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3838
  shows "pi1\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi1\<bullet>\<theta>_n),(pi1\<bullet>\<theta>_c)<(pi1\<bullet>M)>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3839
  and   "pi2\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi2\<bullet>\<theta>_n),(pi2\<bullet>\<theta>_c)<(pi2\<bullet>M)>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3840
apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3841
apply(auto simp add: eq_bij fresh_bij eqvts perm_pi_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3842
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3843
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3844
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3845
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3846
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3847
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3848
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3849
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3850
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3851
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3852
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3853
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3854
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3855
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3856
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3857
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3858
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3859
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3860
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3861
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3862
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3863
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3864
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3865
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3866
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3867
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3868
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3869
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3870
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3871
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3872
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3873
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3874
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3875
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3876
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3877
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3878
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3879
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3880
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3881
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3882
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3883
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3884
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3885
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3886
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3887
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3888
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3889
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3890
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3891
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3892
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3893
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3894
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3895
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3896
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3897
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3898
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3899
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3900
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3901
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3902
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3903
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3904
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3905
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3906
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3907
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3908
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3909
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3910
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3911
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3912
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3913
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3914
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3915
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3916
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3917
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3918
apply(rule case_cong)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3919
apply(rule find_maps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3920
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3921
apply(perm_simp add: eqvts)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3922
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3923
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3924
lemma ax_psubst:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3925
  assumes a: "\<theta>_n,\<theta>_c<M> = Ax x a"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3926
  and     b: "a\<sharp>(\<theta>_n,\<theta>_c)" "x\<sharp>(\<theta>_n,\<theta>_c)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3927
  shows "M = Ax x a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3928
using a b
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3929
apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3930
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3931
apply(drule lookup_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3932
apply(simp)+
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3933
apply(case_tac "findc \<theta>_c coname")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3934
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3935
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3936
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3937
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3938
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3939
apply(case_tac "findn \<theta>_n name")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3940
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3941
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3942
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3943
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3944
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3945
apply(case_tac "findc \<theta>_c coname3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3946
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3947
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3948
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3949
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3950
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3951
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3952
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3953
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3954
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3955
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3956
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3957
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3958
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3959
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3960
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3961
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3962
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3963
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3964
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3965
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3966
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3967
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3968
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3969
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3970
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3971
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3972
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3973
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3974
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3975
apply(case_tac "findn \<theta>_n name3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3976
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3977
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3978
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3979
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3980
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3981
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3982
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3983
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3984
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3985
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3986
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  3987
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3988
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3989
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3990
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3991
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3992
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3993
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3994
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3995
lemma better_Cut_substc1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3996
  assumes a: "a\<sharp>(P,b)" "b\<sharp>N" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3997
  shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.(M{b:=(y).P}) (x).N"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3998
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3999
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4000
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4001
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4002
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4003
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4004
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4005
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4006
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4007
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4008
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4009
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4010
apply(simp add: calc_atm fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4011
apply(subgoal_tac"b\<sharp>([(ca,x)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4012
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4013
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4014
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4015
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4016
apply(simp add: fresh_left calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4017
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4018
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4019
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4020
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4021
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4022
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4023
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4024
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4025
lemma better_Cut_substc2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4026
  assumes a: "x\<sharp>(y,P)" "b\<sharp>(a,M)" "N\<noteq>Ax x b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4027
  shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.M (x).(N{b:=(y).P})"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4028
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4029
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4030
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4031
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4032
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4033
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4034
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4035
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4036
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4037
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4038
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4039
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4040
apply(simp add: calc_atm fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4041
apply(subgoal_tac"b\<sharp>([(c,a)]\<bullet>M)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4042
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4043
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4044
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4045
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4046
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4047
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4048
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4049
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4050
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4051
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4052
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4053
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4054
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4055
lemma better_Cut_substn1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4056
  assumes a: "y\<sharp>(x,N)" "a\<sharp>(b,P)" "M\<noteq>Ax y a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4057
  shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.(M{y:=<b>.P}) (x).N"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4058
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4059
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4060
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4061
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4062
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4063
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4064
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4065
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4066
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4067
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4068
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4069
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4070
apply(simp add: calc_atm fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4071
apply(subgoal_tac"y\<sharp>([(ca,x)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4072
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4073
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4074
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4075
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4076
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4077
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4078
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4079
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4080
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4081
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4082
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4083
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4084
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4085
lemma better_Cut_substn2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4086
  assumes a: "x\<sharp>(P,y)" "y\<sharp>M" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4087
  shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.M (x).(N{y:=<b>.P})"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4088
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4089
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4090
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4091
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4092
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4093
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4094
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4095
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4096
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4097
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4098
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4099
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4100
apply(simp add: calc_atm fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4101
apply(subgoal_tac"y\<sharp>([(c,a)]\<bullet>M)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4102
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4103
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4104
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4105
apply(perm_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4106
apply(simp add: fresh_left calc_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4107
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4108
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4109
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4110
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4111
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4112
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4113
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4114
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4115
lemma psubst_fresh_name:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4116
  fixes x::"name"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4117
  assumes a: "x\<sharp>\<theta>_n" "x\<sharp>\<theta>_c" "x\<sharp>M"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4118
  shows "x\<sharp>\<theta>_n,\<theta>_c<M>"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4119
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4120
apply(nominal_induct M avoiding: x \<theta>_n \<theta>_c rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4121
apply(simp add: lookup_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4122
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4123
apply(simp add: lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4124
apply(simp add: lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4125
apply(simp add: lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4126
apply(simp add: lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4127
apply(simp add: lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4128
apply(simp add: lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4129
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4130
apply(case_tac "findc \<theta>_c coname")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4131
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4132
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4133
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4134
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4135
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4136
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4137
apply(case_tac "findn \<theta>_n name")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4138
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4139
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4140
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4141
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4142
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4143
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4144
apply(case_tac "findc \<theta>_c coname3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4145
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4146
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4147
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4148
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4149
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4150
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4151
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4152
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4153
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4154
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4155
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4156
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4157
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4158
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4159
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4160
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4161
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4162
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4163
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4164
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4165
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4166
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4167
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4168
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4169
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4170
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4171
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4172
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4173
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4174
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4175
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4176
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4177
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4178
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4179
apply(case_tac "findn \<theta>_n name3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4180
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4181
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4182
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4183
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4184
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4185
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4186
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4187
apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4188
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4189
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4190
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4191
apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4192
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4193
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4194
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4195
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4196
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4197
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4198
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4199
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4200
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4201
lemma psubst_fresh_coname:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4202
  fixes a::"coname"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4203
  assumes a: "a\<sharp>\<theta>_n" "a\<sharp>\<theta>_c" "a\<sharp>M"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4204
  shows "a\<sharp>\<theta>_n,\<theta>_c<M>"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4205
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4206
apply(nominal_induct M avoiding: a \<theta>_n \<theta>_c rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4207
apply(simp add: lookup_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4208
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4209
apply(simp add: lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4210
apply(simp add: lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4211
apply(simp add: lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4212
apply(simp add: lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4213
apply(simp add: lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4214
apply(simp add: lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4215
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4216
apply(case_tac "findc \<theta>_c coname")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4217
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4218
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4219
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4220
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4221
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4222
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4223
apply(case_tac "findn \<theta>_n name")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4224
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4225
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4226
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4227
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4228
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4229
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4230
apply(case_tac "findc \<theta>_c coname3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4231
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4232
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4233
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4234
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4235
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4236
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4237
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4238
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4239
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4240
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4241
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4242
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4243
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4244
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4245
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4246
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4247
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4248
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4249
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4250
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4251
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4252
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4253
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4254
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4255
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4256
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4257
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4258
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4259
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4260
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4261
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4262
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4263
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4264
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4265
apply(case_tac "findn \<theta>_n name3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4266
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4267
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4268
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4269
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4270
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4271
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4272
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4273
apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4274
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4275
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4276
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4277
apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4278
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4279
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4280
apply(auto simp add: abs_fresh)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4281
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4282
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4283
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4284
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4285
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4286
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4287
lemma psubst_csubst:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4288
  assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4289
  shows "\<theta>_n,((a,x,P)#\<theta>_c)<M> = ((\<theta>_n,\<theta>_c<M>){a:=(x).P})"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4290
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4291
apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4292
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4293
apply(simp add: lookup_csubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4294
apply(simp add: fresh_list_cons fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4295
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4296
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4297
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4298
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4299
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4300
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4301
apply(simp add: lookupd_fresh)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4302
apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4303
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4304
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4305
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4306
apply(simp add: alpha nrename_swap fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4307
apply(rule lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4308
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4309
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4310
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4311
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4312
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4313
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4314
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4315
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4316
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4317
apply(simp add: lookupd_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4318
apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4319
apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4320
apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4321
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4322
apply(rule lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4323
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4324
apply(rule lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4325
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4326
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4327
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4328
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4329
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4330
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4331
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4332
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4333
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4334
apply(drule ax_psubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4335
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4336
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4337
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4338
apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4339
apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4340
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4341
apply(rule lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4342
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4343
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4344
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4345
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4346
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4347
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4348
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4349
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4350
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4351
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4352
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4353
apply(simp add: alpha)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4354
apply(simp add: alpha nrename_swap fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4355
apply(simp add: lookupd_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4356
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4357
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4358
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4359
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4360
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4361
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4362
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4363
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4364
apply(simp add: lookupd_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4365
apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4366
apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4367
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4368
apply(rule lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4369
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4370
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4371
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4372
apply(rule better_Cut_substc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4373
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4374
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4375
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4376
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4377
apply(drule ax_psubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4378
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4379
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4380
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4381
(* NotR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4382
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4383
apply(case_tac "findc \<theta>_c coname")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4384
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4385
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4386
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4387
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4388
apply(drule cmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4389
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4390
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4391
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4392
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4393
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4394
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4395
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4396
apply(rule better_Cut_substc1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4397
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4398
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4399
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4400
(* NotL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4401
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4402
apply(case_tac "findn \<theta>_n name")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4403
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4404
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4405
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4406
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4407
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4408
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4409
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4410
apply(drule_tac a="a" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4411
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4412
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4413
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4414
apply(rule better_Cut_substc2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4415
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4416
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4417
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4418
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4419
(* AndR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4420
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4421
apply(case_tac "findc \<theta>_c coname3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4422
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4423
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4424
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4425
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4426
apply(drule cmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4427
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4428
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4429
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4430
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4431
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4432
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4433
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4434
apply(rule better_Cut_substc1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4435
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4436
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4437
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4438
(* AndL1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4439
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4440
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4441
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4442
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4443
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4444
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4445
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4446
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4447
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4448
apply(drule_tac a="a" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4449
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4450
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4451
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4452
apply(rule better_Cut_substc2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4453
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4454
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4455
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4456
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4457
(* AndL2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4458
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4459
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4460
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4461
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4462
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4463
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4464
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4465
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4466
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4467
apply(drule_tac a="a" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4468
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4469
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4470
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4471
apply(rule better_Cut_substc2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4472
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4473
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4474
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4475
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4476
(* OrR1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4477
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4478
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4479
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4480
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4481
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4482
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4483
apply(drule cmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4484
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4485
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4486
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4487
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4488
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4489
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4490
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4491
apply(rule better_Cut_substc1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4492
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4493
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4494
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4495
(* OrR2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4496
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4497
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4498
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4499
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4500
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4501
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4502
apply(drule cmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4503
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4504
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4505
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4506
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4507
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4508
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4509
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4510
apply(rule better_Cut_substc1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4511
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4512
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4513
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4514
(* OrL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4515
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4516
apply(case_tac "findn \<theta>_n name3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4517
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4518
apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4519
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4520
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4521
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4522
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4523
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4524
apply(drule_tac a="a" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4525
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4526
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4527
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4528
apply(rule better_Cut_substc2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4529
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4530
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4531
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4532
apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4533
(* ImpR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4534
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4535
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4536
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4537
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4538
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4539
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4540
apply(drule cmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4541
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4542
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4543
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4544
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4545
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4546
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4547
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4548
apply(rule better_Cut_substc1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4549
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4550
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4551
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4552
(* ImpL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4553
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4554
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4555
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4556
apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4557
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4558
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4559
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4560
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4561
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4562
apply(simp add: abs_fresh subst_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4563
apply(drule_tac a="a" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4564
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4565
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4566
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4567
apply(rule better_Cut_substc2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4568
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4569
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4570
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4571
apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4572
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4573
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4574
lemma psubst_nsubst:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4575
  assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4576
  shows "((x,a,P)#\<theta>_n),\<theta>_c<M> = ((\<theta>_n,\<theta>_c<M>){x:=<a>.P})"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4577
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4578
apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4579
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4580
apply(simp add: lookup_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4581
apply(rule lookupb_lookupa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4582
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4583
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4584
apply(rule forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4585
apply(rule lookup_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4586
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4587
apply(auto simp add: lookupc_freshness fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4588
apply(simp add: lookupc_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4589
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4590
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4591
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4592
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4593
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4594
apply(simp add: lookupd_fresh)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4595
apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4596
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4597
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4598
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4599
apply(simp add: alpha crename_swap fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4600
apply(rule lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4601
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4602
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4603
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4604
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4605
apply(simp add: abs_fresh) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4606
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4607
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4608
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4609
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4610
apply(simp add: lookupc_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4611
apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4612
apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4613
apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4614
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4615
apply(rule lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4616
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4617
apply(rule lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4618
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4619
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4620
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4621
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4622
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4623
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4624
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4625
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4626
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4627
apply(simp add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4628
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4629
apply(simp add: alpha crename_swap fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4630
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4631
apply(simp add: lookupc_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4632
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4633
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4634
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4635
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4636
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4637
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4638
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4639
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4640
apply(simp add: lookupc_unicity)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4641
apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4642
apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4643
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4644
apply(rule lookupc_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4645
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4646
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4647
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4648
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4649
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4650
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4651
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4652
apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4653
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4654
apply(drule ax_psubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4655
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4656
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4657
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4658
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4659
apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4660
apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4661
apply(simp add: forget)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4662
apply(rule lookupd_freshness)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4663
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4664
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4665
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4666
apply(rule better_Cut_substn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4667
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4668
apply(simp add: abs_fresh fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4669
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4670
apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4671
apply(drule ax_psubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4672
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4673
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4674
apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4675
(* NotR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4676
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4677
apply(case_tac "findc \<theta>_c coname")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4678
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4679
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4680
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4681
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4682
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4683
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4684
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4685
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4686
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4687
apply(rule better_Cut_substn1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4688
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4689
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4690
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4691
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4692
(* NotL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4693
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4694
apply(case_tac "findn \<theta>_n name")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4695
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4696
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4697
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4698
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4699
apply(drule nmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4700
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4701
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4702
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4703
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4704
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4705
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4706
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4707
apply(rule better_Cut_substn2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4708
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4709
apply(simp add: nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4710
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4711
(* AndR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4712
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4713
apply(case_tac "findc \<theta>_c coname3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4714
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4715
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4716
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4717
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4718
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4719
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4720
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4721
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4722
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4723
apply(rule better_Cut_substn1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4724
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4725
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4726
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4727
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4728
(* AndL1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4729
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4730
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4731
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4732
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4733
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4734
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4735
apply(drule nmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4736
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4737
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4738
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4739
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4740
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4741
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4742
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4743
apply(rule better_Cut_substn2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4744
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4745
apply(simp add: nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4746
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4747
(* AndL2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4748
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4749
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4750
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4751
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4752
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4753
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4754
apply(drule nmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4755
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4756
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4757
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4758
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4759
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4760
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4761
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4762
apply(rule better_Cut_substn2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4763
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4764
apply(simp add: nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4765
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4766
(* OrR1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4767
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4768
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4769
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4770
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4771
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4772
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4773
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4774
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4775
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4776
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4777
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4778
apply(rule better_Cut_substn1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4779
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4780
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4781
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4782
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4783
(* OrR2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4784
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4785
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4786
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4787
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4788
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4789
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4790
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4791
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4792
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4793
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4794
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4795
apply(rule better_Cut_substn1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4796
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4797
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4798
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4799
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4800
(* OrL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4801
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4802
apply(case_tac "findn \<theta>_n name3")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4803
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4804
apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4805
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4806
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4807
apply(drule nmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4808
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4809
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4810
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4811
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4812
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4813
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4814
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4815
apply(rule better_Cut_substn2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4816
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4817
apply(simp add: nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4818
apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4819
(* ImpR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4820
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4821
apply(case_tac "findc \<theta>_c coname2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4822
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4823
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4824
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4825
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4826
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4827
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4828
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4829
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4830
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4831
apply(rule better_Cut_substn1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4832
apply(simp add: cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4833
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4834
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4835
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4836
(* ImpL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4837
apply(simp)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4838
apply(case_tac "findn \<theta>_n name2")
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4839
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4840
apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4841
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4842
apply(auto simp add: fresh_list_cons fresh_prod)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4843
apply(drule nmaps_false)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4844
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4845
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4846
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4847
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4848
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4849
apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4850
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4851
apply(rule better_Cut_substn2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4852
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4853
apply(simp add: nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4854
apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4855
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4856
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4857
definition 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4858
  ncloses :: "(name\<times>coname\<times>trm) list\<Rightarrow>(name\<times>ty) list \<Rightarrow> bool" ("_ ncloses _" [55,55] 55) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4859
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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>)))"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4861
  
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4862
definition 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4863
  ccloses :: "(coname\<times>name\<times>trm) list\<Rightarrow>(coname\<times>ty) list \<Rightarrow> bool" ("_ ccloses _" [55,55] 55) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4864
where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  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>)))"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4866
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4867
lemma ncloses_elim:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4868
  assumes a: "(x,B) \<in> set \<Gamma>"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4869
  and     b: "\<theta>_n ncloses \<Gamma>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4870
  shows "\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4871
using a b by (auto simp add: ncloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4872
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4873
lemma ccloses_elim:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4874
  assumes a: "(a,B) \<in> set \<Delta>"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4875
  and     b: "\<theta>_c ccloses \<Delta>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4876
  shows "\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4877
using a b by (auto simp add: ccloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4878
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4879
lemma ncloses_subset:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4880
  assumes a: "\<theta>_n ncloses \<Gamma>"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4881
  and     b: "set \<Gamma>' \<subseteq> set \<Gamma>"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4882
  shows "\<theta>_n ncloses \<Gamma>'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4883
using a b by (auto  simp add: ncloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4884
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4885
lemma ccloses_subset:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4886
  assumes a: "\<theta>_c ccloses \<Delta>"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4887
  and     b: "set \<Delta>' \<subseteq> set \<Delta>"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4888
  shows "\<theta>_c ccloses \<Delta>'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4889
using a b by (auto  simp add: ccloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4890
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4891
lemma validc_fresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4892
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4893
  and   \<Delta>::"(coname\<times>ty) list"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4894
  assumes a: "a\<sharp>\<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4895
  shows "\<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4896
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4897
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4898
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4899
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4900
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4901
lemma validn_fresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4902
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4903
  and   \<Gamma>::"(name\<times>ty) list"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4904
  assumes a: "x\<sharp>\<Gamma>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4905
  shows "\<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4906
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4907
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4908
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4909
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4910
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4911
lemma ccloses_extend:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4912
  assumes a: "\<theta>_c ccloses \<Delta>" "a\<sharp>\<Delta>" "a\<sharp>\<theta>_c" "(x):P\<in>\<parallel>(B)\<parallel>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4913
  shows "(a,x,P)#\<theta>_c ccloses (a,B)#\<Delta>"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4914
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4915
apply(simp add: ccloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4916
apply(drule validc_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4917
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4918
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4919
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4920
lemma ncloses_extend:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4921
  assumes a: "\<theta>_n ncloses \<Gamma>" "x\<sharp>\<Gamma>" "x\<sharp>\<theta>_n" "<a>:P\<in>\<parallel><B>\<parallel>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  4922
  shows "(x,a,P)#\<theta>_n ncloses (x,B)#\<Gamma>"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4923
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4924
apply(simp add: ncloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4925
apply(drule validn_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4926
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4927
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4928
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4929
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4930
text {* typing relation *}
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4931
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4932
inductive
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4933
   typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4934
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4936
| TNotR:  "\<lbrakk>x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Delta>' = {(a,NOT B)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4937
           \<Longrightarrow> \<Gamma> \<turnstile> NotR (x).M a \<turnstile> \<Delta>'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4938
| TNotL:  "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); set \<Gamma>' = {(x,NOT B)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk>  
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4939
           \<Longrightarrow> \<Gamma>' \<turnstile> NotL <a>.M x \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4941
           \<Longrightarrow> \<Gamma>' \<turnstile> AndL1 (x).M y \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4943
           \<Longrightarrow> \<Gamma>' \<turnstile> AndL2 (x).M y \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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>); 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4945
           set \<Delta>' = {(c,B AND C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4946
           \<Longrightarrow> \<Gamma> \<turnstile> AndR <a>.M <b>.N c \<turnstile> \<Delta>'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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>;
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4948
           set \<Gamma>' = {(z,B OR C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4949
           \<Longrightarrow> \<Gamma>' \<turnstile> OrL (x).M (y).N z \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4951
           \<Longrightarrow> \<Gamma> \<turnstile> OrR1 <a>.M b \<turnstile> \<Delta>'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4953
           \<Longrightarrow> \<Gamma> \<turnstile> OrR2 <a>.M b \<turnstile> \<Delta>'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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>;
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4955
           set \<Gamma>' = {(y,B IMP C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4956
           \<Longrightarrow> \<Gamma>' \<turnstile> ImpL <a>.M (x).N y \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4958
           \<Longrightarrow> \<Gamma> \<turnstile> ImpR (x).<a>.M b \<turnstile> \<Delta>'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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> 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4960
           \<Longrightarrow> \<Gamma> \<turnstile> Cut <a>.M (x).N \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4961
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4962
equivariance typing
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4963
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4964
lemma fresh_set_member:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4965
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4966
  and   a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4967
  shows "x\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> x\<sharp>e"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4968
  and   "a\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> a\<sharp>e"   
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4969
by (induct L) (auto simp add: fresh_list_cons) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4970
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4971
lemma fresh_subset:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4972
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4973
  and   a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4974
  shows "x\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> x\<sharp>L'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4975
  and   "a\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> a\<sharp>L'"   
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4976
apply(induct L' arbitrary: L) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4977
apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4978
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4979
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4980
lemma fresh_subset_ext:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4981
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4982
  and   a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4983
  shows "x\<sharp>L \<Longrightarrow> x\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> x\<sharp>L'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4984
  and   "a\<sharp>L \<Longrightarrow> a\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> a\<sharp>L'"   
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4985
apply(induct L' arbitrary: L) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4986
apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4987
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4988
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4989
lemma fresh_under_insert:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4990
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4991
  and   a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4992
  and   \<Gamma>::"ctxtn"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4993
  and   \<Delta>::"ctxtc"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4994
  shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<noteq>y \<Longrightarrow> set \<Gamma>' = insert (y,B) (set \<Gamma>) \<Longrightarrow> x\<sharp>\<Gamma>'"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4995
  and   "a\<sharp>\<Delta> \<Longrightarrow> a\<noteq>c \<Longrightarrow> set \<Delta>' = insert (c,B) (set \<Delta>) \<Longrightarrow> a\<sharp>\<Delta>'"   
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4996
apply(rule fresh_subset_ext(1))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4997
apply(auto simp add: fresh_prod fresh_atm fresh_ty)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4998
apply(rule fresh_subset_ext(2))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4999
apply(auto simp add: fresh_prod fresh_atm fresh_ty)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5000
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5001
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5002
nominal_inductive typing
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5003
  apply (simp_all add: abs_fresh fresh_atm fresh_list_cons fresh_prod fresh_ty fresh_ctxt 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5004
                       fresh_list_append abs_supp fin_supp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5005
  apply(auto intro: fresh_under_insert)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5006
  done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5007
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5008
lemma validn_elim:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5009
  assumes a: "validn ((x,B)#\<Gamma>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5010
  shows "validn \<Gamma> \<and> x\<sharp>\<Gamma>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5011
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5012
apply(erule_tac validn.cases)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5013
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5014
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5015
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5016
lemma validc_elim:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5017
  assumes a: "validc ((a,B)#\<Delta>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5018
  shows "validc \<Delta> \<and> a\<sharp>\<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5019
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5020
apply(erule_tac validc.cases)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5021
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5022
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5023
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5024
lemma context_fresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5025
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5026
  and   a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5027
  shows "x\<sharp>\<Gamma> \<Longrightarrow> \<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5028
  and   "a\<sharp>\<Delta> \<Longrightarrow> \<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5029
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5030
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5031
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5032
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5033
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5034
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5035
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5036
lemma typing_implies_valid:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5037
  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5038
  shows "validn \<Gamma> \<and> validc \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5039
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5040
apply(nominal_induct rule: typing.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5041
apply(auto dest: validn_elim validc_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5042
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5043
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5044
lemma ty_perm:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5045
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5046
  and   pi2::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5047
  and   B::"ty"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5048
  shows "pi1\<bullet>B=B" and "pi2\<bullet>B=B"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5049
apply(nominal_induct B rule: ty.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5050
apply(auto simp add: perm_string)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5051
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5052
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5053
lemma ctxt_perm:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5054
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5055
  and   pi2::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5056
  and   \<Gamma>::"ctxtn"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5057
  and   \<Delta>::"ctxtc"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5058
  shows "pi2\<bullet>\<Gamma>=\<Gamma>" and "pi1\<bullet>\<Delta>=\<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5059
apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5060
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5061
apply(auto simp add: calc_atm ty_perm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5062
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5063
apply(auto simp add: calc_atm ty_perm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5064
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5065
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5066
lemma typing_Ax_elim1: 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5067
  assumes a: "\<Gamma> \<turnstile> Ax x a \<turnstile> ((a,B)#\<Delta>)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5068
  shows "(x,B)\<in>set \<Gamma>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5069
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5070
apply(erule_tac typing.cases)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5071
apply(simp_all add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5072
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5073
apply(auto dest: validc_elim context_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5074
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5075
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5076
lemma typing_Ax_elim2: 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5077
  assumes a: "((x,B)#\<Gamma>) \<turnstile> Ax x a \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5078
  shows "(a,B)\<in>set \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5079
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5080
apply(erule_tac typing.cases)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5081
apply(simp_all add: trm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5082
apply(auto  dest!: validn_elim context_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5083
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5084
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5085
lemma psubst_Ax_aux: 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5086
  assumes a: "\<theta>_c cmaps a to Some (y,N)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5087
  shows "lookupb x a \<theta>_c c P = Cut <c>.P (y).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5088
using a
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5089
apply(induct \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5090
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5091
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5092
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5093
lemma psubst_Ax:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5094
  assumes a: "\<theta>_n nmaps x to Some (c,P)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5095
  and     b: "\<theta>_c cmaps a to Some (y,N)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5096
  shows "\<theta>_n,\<theta>_c<Ax x a> = Cut <c>.P (y).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5097
using a b
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5098
apply(induct \<theta>_n)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5099
apply(auto simp add: psubst_Ax_aux)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5100
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5101
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5102
lemma psubst_Cut:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5103
  assumes a: "\<forall>x. M\<noteq>Ax x c"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5104
  and     b: "\<forall>a. N\<noteq>Ax x a"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5105
  and     c: "c\<sharp>(\<theta>_n,\<theta>_c,N)" "x\<sharp>(\<theta>_n,\<theta>_c,M)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5106
  shows "\<theta>_n,\<theta>_c<Cut <c>.M (x).N> = Cut <c>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5107
using a b c
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5108
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5109
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5110
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5111
lemma all_CAND: 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5112
  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5113
  and     b: "\<theta>_n ncloses \<Gamma>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5114
  and     c: "\<theta>_c ccloses \<Delta>"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5115
  shows "SNa (\<theta>_n,\<theta>_c<M>)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5116
using a b c
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5117
proof(nominal_induct avoiding: \<theta>_n \<theta>_c rule: typing.strong_induct)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5118
  case (TAx \<Gamma> \<Delta> x B a \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5119
  then show ?case
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5120
    apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5121
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5122
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5123
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5124
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5125
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5126
    apply(erule conjE)+
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5127
    apply(rule_tac s="Cut <c>.P (xa).Pa" and t="\<theta>_n,\<theta>_c<Ax x a>" in subst)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5128
    apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5129
    apply(simp only: psubst_Ax)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5130
    apply(simp add: CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5131
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5132
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5133
  case (TNotR x \<Gamma> B M \<Delta> \<Delta>' a \<theta>_n \<theta>_c) 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5134
  then show ?case 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5135
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5136
    apply(subgoal_tac "(a,NOT B) \<in> set \<Delta>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5137
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5138
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5139
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5140
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5141
    apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5142
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5143
    apply(rule_tac B="NOT B" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5144
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5145
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5146
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5147
    apply(rule_tac x="c" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5148
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5149
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5150
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5151
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5152
    apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5153
    apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5154
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5155
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5156
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5157
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5158
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5159
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5160
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5161
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5162
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5163
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5164
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5165
    apply(simp add: psubst_nsubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5166
    apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5167
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5168
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5169
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5170
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5171
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5172
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5173
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5174
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5175
    apply(rule ccloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5176
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5177
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5178
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5179
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5180
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5181
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5182
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5183
  case (TNotL a \<Delta> \<Gamma> M B \<Gamma>' x \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5184
  then show ?case
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5185
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5186
    apply(subgoal_tac "(x,NOT B) \<in> set \<Gamma>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5187
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5188
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5189
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5190
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5191
    apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5192
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5193
    apply(rule_tac B="NOT B" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5194
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5195
    apply(rule NEG_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5196
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5197
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5198
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5199
    apply(rule_tac x="a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5200
    apply(rule_tac x="ca" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5201
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5202
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5203
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5204
    apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5205
    apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5206
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5207
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5208
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5209
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5210
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5211
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5212
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5213
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5214
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5215
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5216
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5217
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5218
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5219
    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5220
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5221
    apply(rule ncloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5222
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5223
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5224
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5225
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5226
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5227
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5228
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5229
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5230
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5231
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5232
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5233
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5234
  case (TAndL1 x \<Gamma> y B1 M \<Delta> \<Gamma>' B2 \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5235
  then show ?case     
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5236
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5237
    apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5238
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5239
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5240
    apply(erule exE)+ 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5241
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5242
    apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5243
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5244
    apply(rule_tac B="B1 AND B2" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5245
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5246
    apply(rule NEG_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5247
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5248
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5249
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5250
    apply(rule disjI1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5251
    apply(rule_tac x="x" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5252
    apply(rule_tac x="ca" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5253
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5254
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5255
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5256
    apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5257
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5258
    apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5259
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5260
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5261
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5262
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5263
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5264
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5265
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5266
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5267
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5268
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5269
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5270
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5271
    apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5272
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5273
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5274
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5275
    apply(rule ncloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5276
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5277
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5278
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5279
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5280
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5281
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5282
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5283
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5284
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5285
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5286
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5287
  case (TAndL2 x \<Gamma> y B2 M \<Delta> \<Gamma>' B1 \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5288
  then show ?case 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5289
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5290
    apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5291
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5292
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5293
    apply(erule exE)+ 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5294
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5295
    apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5296
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5297
    apply(rule_tac B="B1 AND B2" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5298
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5299
    apply(rule NEG_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5300
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5301
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5302
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5303
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5304
    apply(rule_tac x="x" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5305
    apply(rule_tac x="ca" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5306
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5307
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5308
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5309
    apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5310
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5311
    apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5312
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5313
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5314
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5315
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5316
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5317
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5318
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5319
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5320
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5321
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5322
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5323
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5324
    apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5325
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5326
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5327
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5328
    apply(rule ncloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5329
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5330
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5331
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5332
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5333
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5334
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5335
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5336
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5337
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5338
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5339
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5340
  case (TAndR a \<Delta> N c b M \<Gamma> B C \<Delta>' \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5341
  then show ?case 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5342
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5343
    apply(subgoal_tac "(c,B AND C) \<in> set \<Delta>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5344
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5345
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5346
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5347
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5348
    apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5349
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5350
    apply(rule_tac B="B AND C" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5351
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5352
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5353
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5354
    apply(rule_tac x="ca" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5355
    apply(rule_tac x="a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5356
    apply(rule_tac x="b" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5357
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5358
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5359
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5360
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5361
    apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5362
    apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5363
    apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5364
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5365
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5366
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5367
    apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5368
    apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5369
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5370
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5371
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5372
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5373
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5374
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5375
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5376
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5377
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5378
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5379
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5380
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5381
    apply(simp add: psubst_csubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5382
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5383
    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5384
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5385
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5386
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5387
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5388
    apply(rule ccloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5389
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5390
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5391
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5392
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5393
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5394
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5395
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5396
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5397
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5398
    apply(rule_tac x="b" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5399
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5400
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5401
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5402
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5403
    apply(simp add: psubst_csubst[symmetric])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5404
    apply(rotate_tac 14)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5405
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5406
    apply(drule_tac x="(b,xa,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5407
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5408
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5409
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5410
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5411
    apply(rule ccloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5412
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5413
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5414
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5415
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5416
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5417
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5418
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5419
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5420
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5421
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5422
  case (TOrL x \<Gamma> N z y M B \<Delta> C \<Gamma>' \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5423
  then show ?case 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5424
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5425
    apply(subgoal_tac "(z,B OR C) \<in> set \<Gamma>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5426
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5427
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5428
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5429
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5430
    apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5431
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5432
    apply(rule_tac B="B OR C" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5433
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5434
    apply(rule NEG_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5435
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5436
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5437
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5438
    apply(rule_tac x="x" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5439
    apply(rule_tac x="y" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5440
    apply(rule_tac x="ca" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5441
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5442
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5443
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5444
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5445
    apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5446
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5447
    apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5448
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5449
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5450
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5451
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5452
    apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5453
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5454
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5455
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5456
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5457
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5458
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5459
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5460
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5461
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5462
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5463
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5464
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5465
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5466
    apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5467
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5468
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5469
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5470
    apply(rule ncloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5471
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5472
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5473
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5474
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5475
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5476
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5477
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5478
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5479
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5480
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5481
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5482
    apply(rule_tac x="y" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5483
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5484
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5485
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5486
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5487
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5488
    apply(rotate_tac 14)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5489
    apply(drule_tac x="(y,a,Pa)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5490
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5491
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5492
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5493
    apply(rule ncloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5494
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5495
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5496
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5497
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5498
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5499
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5500
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5501
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5502
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5503
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5504
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5505
  case (TOrR1 a \<Delta> b \<Gamma> M B1 \<Delta>' B2 \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5506
  then show ?case
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5507
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5508
    apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5509
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5510
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5511
    apply(erule exE)+ 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5512
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5513
    apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5514
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5515
    apply(rule_tac B="B1 OR B2" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5516
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5517
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5518
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5519
    apply(rule disjI1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5520
    apply(rule_tac x="a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5521
    apply(rule_tac x="c" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5522
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5523
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5524
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5525
    apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5526
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5527
    apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5528
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5529
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5530
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5531
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5532
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5533
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5534
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5535
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5536
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5537
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5538
    apply(rule impI)    
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5539
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5540
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5541
    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5542
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5543
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5544
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5545
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5546
    apply(rule ccloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5547
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5548
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5549
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5550
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5551
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5552
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5553
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5554
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5555
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5556
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5557
  case (TOrR2 a \<Delta> b \<Gamma> M B2 \<Delta>' B1 \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5558
  then show ?case 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5559
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5560
    apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5561
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5562
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5563
    apply(erule exE)+ 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5564
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5565
    apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5566
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5567
    apply(rule_tac B="B1 OR B2" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5568
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5569
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5570
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5571
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5572
    apply(rule_tac x="a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5573
    apply(rule_tac x="c" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5574
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5575
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5576
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5577
    apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5578
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5579
    apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5580
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5581
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5582
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5583
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5584
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5585
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5586
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5587
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5588
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5589
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5590
    apply(rule impI)    
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5591
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5592
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5593
    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5594
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5595
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5596
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5597
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5598
    apply(rule ccloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5599
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5600
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5601
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5602
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5603
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5604
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5605
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5606
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5607
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5608
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5609
  case (TImpL a \<Delta> N x \<Gamma> M y B C \<Gamma>' \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5610
  then show ?case
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5611
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5612
    apply(subgoal_tac "(y,B IMP C) \<in> set \<Gamma>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5613
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5614
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5615
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5616
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5617
    apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5618
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5619
    apply(rule_tac B="B IMP C" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5620
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5621
    apply(rule NEG_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5622
    apply(simp (no_asm))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5623
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5624
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5625
    apply(rule_tac x="x" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5626
    apply(rule_tac x="a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5627
    apply(rule_tac x="ca" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5628
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5629
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5630
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5631
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5632
    apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5633
    apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5634
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5635
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5636
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5637
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5638
    apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5639
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5640
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5641
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5642
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5643
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5644
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5645
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5646
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5647
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5648
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5649
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5650
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5651
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5652
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5653
    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5654
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5655
    apply(rule ncloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5656
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5657
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5658
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5659
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5660
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5661
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5662
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5663
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5664
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5665
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5666
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5667
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5668
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5669
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5670
    apply(simp del: NEGc.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5671
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5672
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5673
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5674
    apply(rotate_tac 12)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5675
    apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5676
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5677
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5678
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5679
    apply(rule ncloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5680
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5681
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5682
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5683
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5684
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5685
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5686
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5687
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5688
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5689
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5690
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5691
  case (TImpR a \<Delta> b x \<Gamma> B M C \<Delta>' \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5692
  then show ?case
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5693
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5694
    apply(subgoal_tac "(b,B IMP C) \<in> set \<Delta>'")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5695
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5696
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5697
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5698
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5699
    apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5700
    apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5701
    apply(rule_tac B="B IMP C" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5702
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5703
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5704
    apply(rule disjI2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5705
    apply(rule_tac x="x" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5706
    apply(rule_tac x="a" in exI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5707
    apply(rule_tac x="c" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5708
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5709
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5710
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5711
    apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5712
    apply(simp add: abs_fresh fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5713
    apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5714
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5715
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5716
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5717
    apply(rule conjI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5718
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5719
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5720
    apply(simp add: psubst_csubst[symmetric])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5721
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5722
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5723
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5724
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5725
    apply(rule_tac x="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5726
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5727
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5728
    apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5729
    apply(rule_tac t="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>{x:=<aa>.Pb}" and 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5730
                   s="((x,aa,Pb)#\<theta>_n),((a,z,Pa)#\<theta>_c)<M>" in subst)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5731
    apply(rule psubst_nsubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5732
    apply(simp add: fresh_prod fresh_atm fresh_list_cons)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5733
    apply(drule_tac x="(x,aa,Pb)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5734
    apply(drule_tac x="(a,z,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5735
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5736
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5737
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5738
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5739
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5740
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5741
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5742
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5743
    apply(rule ccloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5744
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5745
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5746
    apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5747
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5748
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5749
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5750
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5751
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5752
    apply(simp add: psubst_nsubst[symmetric])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5753
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5754
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5755
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5756
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5757
    apply(rule_tac x="((x,ca,Q)#\<theta>_n),\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5758
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5759
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5760
    apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5761
    apply(rule_tac t="((x,ca,Q)#\<theta>_n),\<theta>_c<M>{a:=(xaa).Pa}" and 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5762
                   s="((x,ca,Q)#\<theta>_n),((a,xaa,Pa)#\<theta>_c)<M>" in subst)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5763
    apply(rule psubst_csubst)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5764
    apply(simp add: fresh_prod fresh_atm fresh_list_cons)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5765
    apply(drule_tac x="(x,ca,Q)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5766
    apply(drule_tac x="(a,xaa,Pa)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5767
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5768
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5769
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5770
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5771
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5772
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5773
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5774
    apply(rule ccloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5775
    apply(rule ccloses_subset)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5776
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5777
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5778
    apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5779
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5780
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5781
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5782
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5783
    apply(blast)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5784
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5785
next
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5786
  case (TCut a \<Delta> N x \<Gamma> M B \<theta>_n \<theta>_c)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5787
  then show ?case 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5788
    apply -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5789
    apply(case_tac "\<forall>y. M\<noteq>Ax y a")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5790
    apply(case_tac "\<forall>c. N\<noteq>Ax x c")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5791
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5792
    apply(rule_tac B="B" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5793
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5794
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5795
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5796
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5797
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5798
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5799
    apply(rule allI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5800
    apply(rule allI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5801
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5802
    apply(simp add: psubst_csubst[symmetric]) (*?*)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5803
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5804
    apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5805
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5806
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5807
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5808
    apply(rule ccloses_extend) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5809
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5810
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5811
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5812
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5813
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5814
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5815
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5816
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5817
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5818
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5819
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5820
    apply(rule allI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5821
    apply(rule allI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5822
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5823
    apply(simp add: psubst_nsubst[symmetric]) (*?*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5824
    apply(rotate_tac 11)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5825
    apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5826
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5827
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5828
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5829
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5830
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5831
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5832
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5833
    apply(drule_tac meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5834
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5835
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5836
    (* cases at least one axiom *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5837
    apply(simp (no_asm_use))
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5838
    apply(erule exE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5839
    apply(simp del: psubst.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5840
    apply(drule typing_Ax_elim2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5841
    apply(auto simp add: trm.inject)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5842
    apply(rule_tac B="B" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5843
    (* left term *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5844
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5845
    apply(unfold BINDINGc_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5846
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5847
    apply(rule_tac x="a" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5848
    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5849
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5850
    apply(rule allI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5851
    apply(rule impI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5852
    apply(drule_tac x="\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5853
    apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5854
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5855
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5856
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5857
    apply(rule ccloses_extend) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5858
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5859
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5860
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5861
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5862
    apply(simp add: psubst_csubst[symmetric]) (*?*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5863
    (* right term -axiom *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5864
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5865
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5866
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5867
    apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5868
    apply(frule_tac y="x" in lookupd_cmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5869
    apply(drule cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5870
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5871
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5872
    apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5873
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5874
    apply(simp add: ntrm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5875
    apply(simp add: alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5876
    apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5877
    apply(rule nrename_swap)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5878
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5879
    (* M is axiom *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5880
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5881
    apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5882
    (* both are axioms *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5883
    apply(rule_tac B="B" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5884
    apply(drule typing_Ax_elim1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5885
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5886
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5887
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5888
    apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5889
    apply(frule_tac a="a" in lookupc_nmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5890
    apply(drule_tac a="a" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5891
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5892
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5893
    apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5894
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5895
    apply(simp add: ctrm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5896
    apply(simp add: alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5897
    apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5898
    apply(rule crename_swap)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5899
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5900
    apply(drule typing_Ax_elim2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5901
    apply(drule ccloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5902
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5903
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5904
    apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5905
    apply(frule_tac y="x" in lookupd_cmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5906
    apply(drule cmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5907
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5908
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5909
    apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5910
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5911
    apply(simp add: ntrm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5912
    apply(simp add: alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5913
    apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5914
    apply(rule nrename_swap)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5915
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5916
    (* N is not axioms *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5917
    apply(rule_tac B="B" in CUT_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5918
    (* left term *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5919
    apply(drule typing_Ax_elim1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5920
    apply(drule ncloses_elim)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5921
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5922
    apply(erule exE)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5923
    apply(erule conjE)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5924
    apply(frule_tac a="a" in lookupc_nmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5925
    apply(drule_tac a="a" in nmaps_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5926
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5927
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5928
    apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5929
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5930
    apply(simp add: ctrm.inject)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5931
    apply(simp add: alpha fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5932
    apply(rule sym)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5933
    apply(rule crename_swap)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5934
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5935
    apply(rule BINDING_implies_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5936
    apply(unfold BINDINGn_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5937
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5938
    apply(rule_tac x="x" in exI)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5939
    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5940
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5941
    apply(rule allI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5942
    apply(rule allI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5943
    apply(rule impI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5944
    apply(simp add: psubst_nsubst[symmetric]) (*?*)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5945
    apply(rotate_tac 10)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5946
    apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  5947
    apply(drule_tac x="\<theta>_c" in meta_spec)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5948
    apply(drule meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5949
    apply(rule ncloses_extend)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5950
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5951
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5952
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5953
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5954
    apply(drule_tac meta_mp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5955
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5956
    apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5957
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5958
qed
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5959
39246
9e58f0499f57 modernized primrec
haftmann
parents: 36277
diff changeset
  5960
primrec "idn" :: "(name\<times>ty) list\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list" where
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5961
  "idn [] a   = []"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 36277
diff changeset
  5962
| "idn (p#\<Gamma>) a = ((fst p),a,Ax (fst p) a)#(idn \<Gamma> a)"
9e58f0499f57 modernized primrec
haftmann
parents: 36277
diff changeset
  5963
9e58f0499f57 modernized primrec
haftmann
parents: 36277
diff changeset
  5964
primrec "idc" :: "(coname\<times>ty) list\<Rightarrow>name\<Rightarrow>(coname\<times>name\<times>trm) list" where
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5965
  "idc [] x    = []"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 36277
diff changeset
  5966
| "idc (p#\<Delta>) x = ((fst p),x,Ax x (fst p))#(idc \<Delta> x)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5967
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5968
lemma idn_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5969
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5970
  and   pi2::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5971
  shows "(pi1\<bullet>(idn \<Gamma> a)) = idn (pi1\<bullet>\<Gamma>) (pi1\<bullet>a)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5972
  and   "(pi2\<bullet>(idn \<Gamma> a)) = idn (pi2\<bullet>\<Gamma>) (pi2\<bullet>a)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5973
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5974
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5975
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5976
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5977
lemma idc_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5978
  fixes pi1::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5979
  and   pi2::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5980
  shows "(pi1\<bullet>(idc \<Delta> x)) = idc (pi1\<bullet>\<Delta>) (pi1\<bullet>x)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5981
  and   "(pi2\<bullet>(idc \<Delta> x)) = idc (pi2\<bullet>\<Delta>) (pi2\<bullet>x)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5982
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5983
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5984
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5985
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5986
lemma ccloses_id:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5987
  shows "(idc \<Delta> x) ccloses \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5988
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5989
apply(auto simp add: ccloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5990
apply(rule Ax_in_CANDs)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5991
apply(rule Ax_in_CANDs)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5992
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5993
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5994
lemma ncloses_id:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5995
  shows "(idn \<Gamma> a) ncloses \<Gamma>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5996
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5997
apply(auto simp add: ncloses_def)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5998
apply(rule Ax_in_CANDs)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5999
apply(rule Ax_in_CANDs)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6000
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6001
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6002
lemma fresh_idn:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6003
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6004
  and   a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6005
  shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<sharp>idn \<Gamma> a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6006
  and   "a\<sharp>(\<Gamma>,b) \<Longrightarrow> a\<sharp>idn \<Gamma> b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6007
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6008
apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6009
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6010
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6011
lemma fresh_idc:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6012
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6013
  and   a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6014
  shows "x\<sharp>(\<Delta>,y) \<Longrightarrow> x\<sharp>idc \<Delta> y"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6015
  and   "a\<sharp>\<Delta>  \<Longrightarrow> a\<sharp>idc \<Delta> y"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6016
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6017
apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6018
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6019
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6020
lemma idc_cmaps:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6021
  assumes a: "idc \<Delta> y cmaps b to Some (x,M)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6022
  shows "M=Ax x b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6023
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6024
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6025
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6026
apply(case_tac "b=a")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6027
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6028
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6029
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6030
lemma idn_nmaps:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6031
  assumes a: "idn \<Gamma> a nmaps x to Some (b,M)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6032
  shows "M=Ax x b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6033
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6034
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6035
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6036
apply(case_tac "aa=x")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6037
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6038
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6039
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6040
lemma lookup1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6041
  assumes a: "x\<sharp>(idn \<Gamma> b)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  6042
  shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupa x a \<theta>_c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6043
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6044
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6045
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6046
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6047
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6048
lemma lookup2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6049
  assumes a: "\<not>(x\<sharp>(idn \<Gamma> b))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 39246
diff changeset
  6050
  shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupb x a \<theta>_c b (Ax x b)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6051
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6052
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6053
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6054
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6055
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6056
lemma lookup3:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6057
  assumes a: "a\<sharp>(idc \<Delta> y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6058
  shows "lookupa x a (idc \<Delta> y) = Ax x a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6059
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6060
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6061
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6062
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6063
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6064
lemma lookup4:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6065
  assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6066
  shows "lookupa x a (idc \<Delta> y) = Cut <a>.(Ax x a) (y).Ax y a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6067
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6068
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6069
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6070
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6071
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6072
lemma lookup5:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6073
  assumes a: "a\<sharp>(idc \<Delta> y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6074
  shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (x).Ax x a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6075
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6076
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6077
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6078
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6079
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6080
lemma lookup6:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6081
  assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6082
  shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (y).Ax y a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6083
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6084
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6085
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6086
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6087
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6088
lemma lookup7:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6089
  shows "lookupc x a (idn \<Gamma> b) = Ax x a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6090
apply(induct \<Gamma>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6091
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6092
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6093
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6094
lemma lookup8:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6095
  shows "lookupd x a (idc \<Delta> y) = Ax x a"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6096
apply(induct \<Delta>)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6097
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6098
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6099
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6100
lemma id_redu:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  6101
  shows "(idn \<Gamma> x),(idc \<Delta> a)<M> \<longrightarrow>\<^sub>a* M"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6102
apply(nominal_induct M avoiding: \<Gamma> \<Delta> x a rule: trm.strong_induct)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6103
apply(auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6104
(* Ax *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6105
apply(case_tac "name\<sharp>(idn \<Gamma> x)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6106
apply(simp add: lookup1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6107
apply(case_tac "coname\<sharp>(idc \<Delta> a)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6108
apply(simp add: lookup3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6109
apply(simp add: lookup4)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6110
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6111
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6112
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6113
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6114
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6115
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6116
apply(simp add: lookup2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6117
apply(case_tac "coname\<sharp>(idc \<Delta> a)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6118
apply(simp add: lookup5)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6119
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6120
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6121
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6122
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6123
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6124
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6125
apply(simp add: lookup6)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6126
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6127
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6128
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6129
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6130
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6131
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6132
(* Cut *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6133
apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name psubst_fresh_coname fresh_atm fresh_prod )[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6134
apply(simp add: lookup7 lookup8)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6135
apply(simp add: lookup7 lookup8)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6136
apply(simp add: a_star_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6137
apply(simp add: lookup7 lookup8)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6138
apply(simp add: a_star_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6139
apply(simp add: a_star_Cut)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6140
(* NotR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6141
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6142
apply(case_tac "findc (idc \<Delta> a) coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6143
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6144
apply(simp add: a_star_NotR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6145
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6146
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6147
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6148
apply(drule idc_cmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6149
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6150
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6151
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6152
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6153
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6154
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6155
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6156
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6157
apply(simp add: crename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6158
apply(simp add: a_star_NotR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6159
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6160
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6161
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6162
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6163
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6164
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6165
(* NotL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6166
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6167
apply(case_tac "findn (idn \<Gamma> x) name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6168
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6169
apply(simp add: a_star_NotL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6170
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6171
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6172
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6173
apply(drule idn_nmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6174
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6175
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6176
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6177
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6178
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6179
apply(rule better_LAxL_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6180
apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6181
apply(assumption)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6182
apply(simp add: nrename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6183
apply(simp add: a_star_NotL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6184
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6185
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6186
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6187
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6188
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6189
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6190
(* AndR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6191
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6192
apply(case_tac "findc (idc \<Delta> a) coname3")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6193
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6194
apply(simp add: a_star_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6195
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6196
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6197
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6198
apply(drule idc_cmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6199
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6200
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6201
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6202
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6203
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6204
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6205
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6206
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6207
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6208
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6209
apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name crename_fresh fresh_atm fresh_prod )[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6210
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6211
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6212
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6213
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6214
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6215
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6216
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6217
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6218
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6219
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6220
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6221
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6222
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6223
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6224
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6225
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6226
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6227
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6228
apply(simp add: crename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6229
apply(simp add: a_star_AndR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6230
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6231
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6232
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6233
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6234
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6235
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6236
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6237
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6238
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6239
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6240
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6241
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6242
(* AndL1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6243
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6244
apply(case_tac "findn (idn \<Gamma> x) name2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6245
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6246
apply(simp add: a_star_AndL1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6247
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6248
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6249
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6250
apply(drule idn_nmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6251
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6252
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6253
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6254
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6255
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6256
apply(rule better_LAxL_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6257
apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6258
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6259
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6260
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6261
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6262
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6263
apply(simp add: nrename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6264
apply(simp add: a_star_AndL1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6265
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6266
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6267
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6268
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6269
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6270
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6271
(* AndL2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6272
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6273
apply(case_tac "findn (idn \<Gamma> x) name2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6274
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6275
apply(simp add: a_star_AndL2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6276
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6277
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6278
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6279
apply(drule idn_nmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6280
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6281
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6282
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6283
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6284
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6285
apply(rule better_LAxL_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6286
apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6287
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6288
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6289
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6290
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6291
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6292
apply(simp add: nrename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6293
apply(simp add: a_star_AndL2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6294
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6295
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6296
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6297
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6298
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6299
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6300
(* OrR1 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6301
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6302
apply(case_tac "findc (idc \<Delta> a) coname2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6303
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6304
apply(simp add: a_star_OrR1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6305
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6306
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6307
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6308
apply(drule idc_cmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6309
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6310
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6311
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6312
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6313
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6314
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6315
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6316
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6317
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6318
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6319
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6320
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6321
apply(simp add: crename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6322
apply(simp add: a_star_OrR1)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6323
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6324
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6325
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6326
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6327
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6328
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6329
(* OrR2 *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6330
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6331
apply(case_tac "findc (idc \<Delta> a) coname2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6332
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6333
apply(simp add: a_star_OrR2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6334
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6335
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6336
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6337
apply(drule idc_cmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6338
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6339
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6340
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6341
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6342
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6343
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6344
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6345
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6346
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6347
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6348
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6349
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6350
apply(simp add: crename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6351
apply(simp add: a_star_OrR2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6352
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6353
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6354
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6355
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6356
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6357
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6358
(* OrL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6359
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6360
apply(case_tac "findn (idn \<Gamma> x) name3")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6361
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6362
apply(simp add: a_star_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6363
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6364
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6365
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6366
apply(drule idn_nmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6367
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6368
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6369
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6370
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6371
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6372
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6373
apply(rule better_LAxL_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6374
apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6375
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6376
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6377
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6378
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6379
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6380
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6381
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6382
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6383
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6384
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6385
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6386
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6387
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6388
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6389
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6390
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6391
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6392
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6393
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6394
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6395
apply(simp add: nrename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6396
apply(simp add: a_star_OrL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6397
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6398
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6399
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6400
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6401
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6402
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6403
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6404
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6405
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6406
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6407
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6408
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6409
(* ImpR *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6410
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6411
apply(case_tac "findc (idc \<Delta> a) coname2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6412
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6413
apply(simp add: a_star_ImpR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6414
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6415
apply(generate_fresh "coname")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6416
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6417
apply(drule idc_cmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6418
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6419
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6420
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6421
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6422
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6423
apply(rule better_LAxR_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6424
apply(rule fic.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6425
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6426
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6427
apply(rule crename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6428
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6429
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6430
apply(simp add: crename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6431
apply(simp add: a_star_ImpR)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6432
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6433
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6434
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6435
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6436
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6437
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6438
(* ImpL *)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6439
apply(simp add: fresh_idn fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6440
apply(case_tac "findn (idn \<Gamma> x) name2")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6441
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6442
apply(simp add: a_star_ImpL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6443
apply(auto)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6444
apply(generate_fresh "name")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6445
apply(fresh_fun_simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6446
apply(drule idn_nmaps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6447
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6448
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6449
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6450
apply(rule a_star_trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6451
apply(rule a_starI)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6452
apply(rule al_redu)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6453
apply(rule better_LAxL_intro)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6454
apply(rule fin.intros)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6455
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6456
apply(simp add: abs_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6457
apply(rule aux3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6458
apply(rule nrename.simps)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6459
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6460
apply(rule psubst_fresh_coname)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6461
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6462
apply(simp add: fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6463
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6464
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6465
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6466
apply(auto simp add: fresh_prod fresh_atm)[1]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6467
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6468
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6469
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6470
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6471
apply(simp add: fresh_prod fresh_atm)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6472
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6473
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6474
apply(simp add: nrename_fresh)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6475
apply(simp add: a_star_ImpL)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6476
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6477
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6478
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6479
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6480
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6481
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6482
apply(rule psubst_fresh_name)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6483
apply(rule fresh_idn)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6484
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6485
apply(rule fresh_idc)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6486
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6487
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6488
done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6489
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6490
theorem ALL_SNa:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6491
  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6492
  shows "SNa M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6493
proof -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6494
  fix x have "(idc \<Delta> x) ccloses \<Delta>" by (simp add: ccloses_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6495
  moreover
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6496
  fix a have "(idn \<Gamma> a) ncloses \<Gamma>" by (simp add: ncloses_id)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6497
  ultimately have "SNa ((idn \<Gamma> a),(idc \<Delta> x)<M>)" using a by (simp add: all_CAND)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6498
  moreover
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50252
diff changeset
  6499
  have "((idn \<Gamma> a),(idc \<Delta> x)<M>) \<longrightarrow>\<^sub>a* M" by (simp add: id_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6500
  ultimately show "SNa M" by (simp add: a_star_preserves_SNa)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6501
qed
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6502
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6503
end