src/HOL/Nominal/Examples/Class1.thy
author paulson <lp15@cam.ac.uk>
Sat, 20 Apr 2024 23:02:47 +0100
changeset 80139 fec5a23017b5
parent 80138 a30a1385f7d0
child 80172 6c62605cb3f6
permissions -rw-r--r--
Tidying up more messy proofs
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 Class1
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
     2
imports "HOL-Nominal.Nominal"
36277
9be4ab2acc13 split Class.thy into parts to 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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
     5
section \<open>Term-Calculus from Urban's PhD\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
atom_decl name 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
     8
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
     9
text \<open>types\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    10
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73932
diff changeset
    11
no_notation not  ("NOT")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73932
diff changeset
    12
36277
9be4ab2acc13 split Class.thy into parts to 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
nominal_datatype 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
    14
    PR "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
    15
  | NOT  "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
    16
  | AND  "ty" "ty"   ("_ AND _" [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
    17
  | OR   "ty" "ty"   ("_ OR _" [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
    18
  | IMP  "ty" "ty"   ("_ IMP _" [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
    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
    20
instantiation ty :: size
9be4ab2acc13 split Class.thy into parts to 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
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
    22
9be4ab2acc13 split Class.thy into parts to 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
nominal_primrec size_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
    24
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
    25
  "size (PR s)     = (1::nat)"
9be4ab2acc13 split Class.thy into parts to 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
| "size (NOT T)     = 1 + size T"
9be4ab2acc13 split Class.thy into parts to 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
| "size (T1 AND T2) = 1 + size T1 + size T2"
9be4ab2acc13 split Class.thy into parts to 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
| "size (T1 OR T2)  = 1 + size T1 + size T2"
9be4ab2acc13 split Class.thy into parts to 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
| "size (T1 IMP T2) = 1 + size T1 + size T2"
9be4ab2acc13 split Class.thy into parts to 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
by (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
    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
instance ..
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    33
9be4ab2acc13 split Class.thy into parts to 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
end
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma ty_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
    37
  fixes T::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
    38
  shows "(\<exists>s. T=PR s) \<or> (\<exists>T'. T=NOT T') \<or> (\<exists>S U. T=S OR U) \<or> (\<exists>S U. T=S AND U) \<or> (\<exists>S U. T=S IMP 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
    39
by (induct T rule:ty.induct) (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
    40
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
    42
  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
    43
  and   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
    44
  and   T::"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
    45
  shows "a\<sharp>T" and "x\<sharp>T"
9be4ab2acc13 split Class.thy into parts to 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
by (nominal_induct T rule: ty.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    47
   (auto simp: fresh_string)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    48
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    49
text \<open>terms\<close>                               
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
nominal_datatype 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
    52
    Ax   "name" "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
    53
  | Cut  "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"            ("Cut <_>._ (_)._" [100,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
    54
  | NotR "\<guillemotleft>name\<guillemotright>trm" "coname"                 ("NotR (_)._ _" [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
    55
  | NotL "\<guillemotleft>coname\<guillemotright>trm" "name"                 ("NotL <_>._ _" [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
    56
  | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [100,100,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
    57
  | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL1 (_)._ _" [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
    58
  | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL2 (_)._ _" [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
    59
  | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR1 <_>._ _" [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
    60
  | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR2 <_>._ _" [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
    61
  | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"        ("OrL (_)._ (_)._ _" [100,100,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
    62
  | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"       ("ImpR (_).<_>._ _" [100,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
    63
  | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"     ("ImpL <_>._ (_)._ _" [100,100,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
    64
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
    65
text \<open>named terms\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("((_):_)" [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
    68
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
    69
text \<open>conamed terms\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" ("(<_>:_)" [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
    72
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
    73
text \<open>renaming functions\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
nominal_primrec (freshness_context: "(d::coname,e::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
    76
  crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm"  ("_[_\<turnstile>c>_]" [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
    77
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
    78
  "(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else 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
    79
| "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>M\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>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
    80
| "(NotR (x).M a)[d\<turnstile>c>e] = (if a=d then NotR (x).(M[d\<turnstile>c>e]) e else NotR (x).(M[d\<turnstile>c>e]) 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
    81
| "a\<sharp>(d,e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) 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
    82
| "\<lbrakk>a\<sharp>(d,e,N,c);b\<sharp>(d,e,M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>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
    83
          (if c=d then AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d \<turnstile>c>e]) e else AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d\<turnstile>c>e]) 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
    84
| "x\<sharp>y \<Longrightarrow> (AndL1 (x).M y)[d\<turnstile>c>e] = AndL1 (x).(M[d\<turnstile>c>e]) 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
| "x\<sharp>y \<Longrightarrow> (AndL2 (x).M y)[d\<turnstile>c>e] = AndL2 (x).(M[d\<turnstile>c>e]) 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
| "a\<sharp>(d,e,b) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) 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
    87
| "a\<sharp>(d,e,b) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) 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
    88
| "\<lbrakk>x\<sharp>(N,z);y\<sharp>(M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N z)[d\<turnstile>c>e] = OrL (x).(M[d\<turnstile>c>e]) (y).(N[d\<turnstile>c>e]) 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
    89
| "a\<sharp>(d,e,b) \<Longrightarrow> (ImpR (x).<a>.M b)[d\<turnstile>c>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
    90
       (if b=d then ImpR (x).<a>.(M[d\<turnstile>c>e]) e else ImpR (x).<a>.(M[d\<turnstile>c>e]) 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
    91
| "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>(M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y)[d\<turnstile>c>e] = ImpL <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>e]) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    92
  by(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
nominal_primrec (freshness_context: "(u::name,v::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
    95
  nrename :: "trm \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm"      ("_[_\<turnstile>n>_]" [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
    96
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
    97
  "(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else 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
    98
| "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>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
    99
| "x\<sharp>(u,v) \<Longrightarrow> (NotR (x).M a)[u\<turnstile>n>v] = NotR (x).(M[u\<turnstile>n>v]) 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
   100
| "(NotL <a>.M x)[u\<turnstile>n>v] = (if x=u then NotL <a>.(M[u\<turnstile>n>v]) v else NotL <a>.(M[u\<turnstile>n>v]) 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
   101
| "\<lbrakk>a\<sharp>(N,c);b\<sharp>(M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[u\<turnstile>n>v] = AndR <a>.(M[u\<turnstile>n>v]) <b>.(N[u\<turnstile>n>v]) 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
   102
| "x\<sharp>(u,v,y) \<Longrightarrow> (AndL1 (x).M y)[u\<turnstile>n>v] = (if y=u then AndL1 (x).(M[u\<turnstile>n>v]) v else AndL1 (x).(M[u\<turnstile>n>v]) 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
   103
| "x\<sharp>(u,v,y) \<Longrightarrow> (AndL2 (x).M y)[u\<turnstile>n>v] = (if y=u then AndL2 (x).(M[u\<turnstile>n>v]) v else AndL2 (x).(M[u\<turnstile>n>v]) 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
   104
| "a\<sharp>b \<Longrightarrow> (OrR1 <a>.M b)[u\<turnstile>n>v] = OrR1 <a>.(M[u\<turnstile>n>v]) 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
   105
| "a\<sharp>b \<Longrightarrow> (OrR2 <a>.M b)[u\<turnstile>n>v] = OrR2 <a>.(M[u\<turnstile>n>v]) 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
   106
| "\<lbrakk>x\<sharp>(u,v,N,z);y\<sharp>(u,v,M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N z)[u\<turnstile>n>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
   107
        (if z=u then OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) v else OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) 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
   108
| "\<lbrakk>a\<sharp>b; x\<sharp>(u,v)\<rbrakk> \<Longrightarrow> (ImpR (x).<a>.M b)[u\<turnstile>n>v] = ImpR (x).<a>.(M[u\<turnstile>n>v]) 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
   109
| "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y)[u\<turnstile>n>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
   110
        (if y=u then ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) v else ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) y)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   111
  by(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemmas eq_bij = pt_bij[OF pt_name_inst, OF at_name_inst] pt_bij[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
   114
9be4ab2acc13 split Class.thy into parts to 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
lemma crename_name_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
   116
  fixes pi::"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
   117
  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   118
  by (nominal_induct M avoiding: d e rule: trm.strong_induct) (auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma crename_coname_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
   121
  fixes pi::"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
   122
  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   123
  by (nominal_induct M avoiding: d e rule: trm.strong_induct)(auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma nrename_name_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
   126
  fixes pi::"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
   127
  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   128
  by(nominal_induct M avoiding: x y rule: trm.strong_induct) (auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma nrename_coname_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
   131
  fixes pi::"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
   132
  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   133
  by(nominal_induct M avoiding: x y rule: trm.strong_induct)(auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemmas rename_eqvts = crename_name_eqvt crename_coname_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
   136
                      nrename_name_eqvt nrename_coname_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
   137
lemma 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
   138
  assumes a: "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
   139
  shows "M[x\<turnstile>n>y] = 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
   140
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
   141
by (nominal_induct M avoiding: x y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   142
   (auto simp: trm.inject fresh_atm abs_fresh fin_supp abs_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
   145
  assumes a: "a\<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
   146
  shows "M[a\<turnstile>c>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
   147
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
   148
by (nominal_induct M avoiding: a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   149
   (auto simp: trm.inject fresh_atm abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma nrename_nfresh:
9be4ab2acc13 split Class.thy into parts to 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
  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
   153
  shows "x\<sharp>y\<Longrightarrow>x\<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
   154
by (nominal_induct M avoiding: x y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   155
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
 lemma crename_nfresh:
9be4ab2acc13 split Class.thy into parts to 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
  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
   159
  shows "x\<sharp>M\<Longrightarrow>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
   160
by (nominal_induct M avoiding: a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   161
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma crename_cfresh:
9be4ab2acc13 split Class.thy into parts to 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
  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
   165
  shows "a\<sharp>b\<Longrightarrow>a\<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
   166
by (nominal_induct M avoiding: a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   167
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma nrename_cfresh:
9be4ab2acc13 split Class.thy into parts to 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
  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
   171
  shows "c\<sharp>M\<Longrightarrow>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
   172
by (nominal_induct M avoiding: x y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   173
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma nrename_nfresh':
9be4ab2acc13 split Class.thy into parts to 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
  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
   177
  shows "x\<sharp>(M,z,y)\<Longrightarrow>x\<sharp>M[z\<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
   178
by (nominal_induct M avoiding: x z y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   179
   (auto simp: fresh_prod fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma crename_cfresh':
9be4ab2acc13 split Class.thy into parts to 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
  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
   183
  shows "a\<sharp>(M,b,c)\<Longrightarrow>a\<sharp>M[b\<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
   184
by (nominal_induct M avoiding: a b c rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   185
   (auto simp: fresh_prod fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma nrename_rename:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "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
   189
  shows "([(x',x)]\<bullet>M)[x'\<turnstile>n>y]= 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
   190
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
   191
apply(nominal_induct M avoiding: x x' y rule: trm.strong_induct)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   192
 apply(auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
   194
9be4ab2acc13 split Class.thy into parts to 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
lemma crename_rename:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "a'\<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
   197
  shows "([(a',a)]\<bullet>M)[a'\<turnstile>c>b]= 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
   198
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
   199
apply(nominal_induct M avoiding: a a' b rule: trm.strong_induct)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   200
 apply(auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
   202
9be4ab2acc13 split Class.thy into parts to 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
lemmas rename_fresh = nrename_fresh 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
   204
                      nrename_nfresh crename_nfresh crename_cfresh nrename_cfresh
9be4ab2acc13 split Class.thy into parts to 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
                      nrename_nfresh' crename_cfresh'
9be4ab2acc13 split Class.thy into parts to 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
                      nrename_rename crename_rename
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma better_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
   209
  assumes a: "x\<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
   210
  shows "(Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>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
   211
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
   212
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,a,x,u,v)" by (rule exists_fresh(1), rule fin_supp, 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
   213
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,x,u,v)" by (rule exists_fresh(2), rule fin_supp, 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
   214
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   215
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))[u\<turnstile>n>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
   217
                        Cut <a'>.(([(a',a)]\<bullet>M)[u\<turnstile>n>v]) (x').(([(x',x)]\<bullet>N)[u\<turnstile>n>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
   218
    using fs1 fs2
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   219
    by (intro nrename.simps; simp add: fresh_left calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])" using fs1 fs2 a
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   221
    by(simp add: trm.inject alpha fresh_prod rename_eqvts calc_atm rename_fresh fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  finally show "(Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])" using eq1
9be4ab2acc13 split Class.thy into parts to 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
    by simp
9be4ab2acc13 split Class.thy into parts to 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
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
   225
9be4ab2acc13 split Class.thy into parts to 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
lemma better_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
   227
  assumes a: "a\<sharp>(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
   228
  shows "(Cut <a>.M (x).N)[b\<turnstile>c>c] = Cut <a>.(M[b\<turnstile>c>c]) (x).(N[b\<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
   229
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
   230
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,a,x,b,c)" by (rule exists_fresh(1), rule fin_supp, 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
   231
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,x,b,c)" by (rule exists_fresh(2), rule fin_supp, 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
   232
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   233
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))[b\<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
   235
                        Cut <a'>.(([(a',a)]\<bullet>M)[b\<turnstile>c>c]) (x').(([(x',x)]\<bullet>N)[b\<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
   236
    using fs1 fs2
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   237
    by (intro crename.simps; simp add: fresh_left calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = Cut <a>.(M[b\<turnstile>c>c]) (x).(N[b\<turnstile>c>c])" using fs1 fs2 a
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   239
    by(simp add: trm.inject alpha calc_atm rename_fresh fresh_atm fresh_prod rename_eqvts)
36277
9be4ab2acc13 split Class.thy into parts to 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
  finally show "(Cut <a>.M (x).N)[b\<turnstile>c>c] = Cut <a>.(M[b\<turnstile>c>c]) (x).(N[b\<turnstile>c>c])" using eq1
9be4ab2acc13 split Class.thy into parts to 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
    by simp
9be4ab2acc13 split Class.thy into parts to 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
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
   243
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
   245
  shows "M[a\<turnstile>c>a] = 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
   246
by (nominal_induct M avoiding: a rule: trm.strong_induct) (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
   247
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
   249
  shows "M[x\<turnstile>n>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
   250
by (nominal_induct M avoiding: x rule: trm.strong_induct) (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
   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 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
   253
  shows "x\<sharp>M \<Longrightarrow> [(x,y)]\<bullet>M = M[y\<turnstile>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
   254
by (nominal_induct M avoiding: x y rule: trm.strong_induct) 
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   255
   (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   256
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
   259
  shows "a\<sharp>M \<Longrightarrow> [(a,b)]\<bullet>M = M[b\<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
   260
by (nominal_induct M avoiding: a b rule: trm.strong_induct) 
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   261
   (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   262
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma crename_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
   265
  assumes a: "M[a\<turnstile>c>b] = Ax x c" "c\<noteq>a" "c\<noteq>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
   266
  shows "M = 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
   267
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
   268
apply(nominal_induct M avoiding: a b x 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
   269
apply(simp_all add: trm.inject 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
   270
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
   271
9be4ab2acc13 split Class.thy into parts to 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
lemma nrename_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
   273
  assumes a: "M[x\<turnstile>n>y] = Ax z a" "z\<noteq>x" "z\<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
   274
  shows "M = Ax z 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
   275
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
   276
apply(nominal_induct M avoiding: x y z 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
   277
apply(simp_all add: trm.inject 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
   278
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
   279
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
   280
text \<open>substitution functions\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_perm_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
   283
  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
   284
  and   pi::"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
   285
  and   M::"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
   286
  assumes a: "c\<sharp>pi" "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
   287
  shows "c\<sharp>(pi\<bullet>M)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   288
  by (simp add: assms fresh_perm_app)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_perm_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
   291
  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
   292
  and   pi::"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
   293
  and   M::"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
   294
  assumes a: "x\<sharp>pi" "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
   295
  shows "x\<sharp>(pi\<bullet>M)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   296
  by (simp add: assms fresh_perm_app)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   299
  assumes a: "x'\<sharp>P" "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
   300
  shows "fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x') = Cut <c>.P (x').NotL <a>.M x'"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   301
proof (rule fresh_fun_app)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   302
  show "pt (TYPE(trm)::trm itself) (TYPE(name)::name itself)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   303
    by(rule pt_name_inst)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   304
  show "at (TYPE(name)::name itself)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   305
    by(rule at_name_inst)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   306
  show "finite (supp (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')::name set)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   307
    using a by(finite_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   308
  obtain n::name where n: "n\<sharp>(c,P,a,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   309
    by (metis assms fresh_atm(3) fresh_prod)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   310
  with assms
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   311
  show "\<exists>b. b \<sharp> (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x', Cut <c>.P (b).NotL <a>.M b)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   312
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   313
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   314
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   315
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   316
  show "x' \<sharp> (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   317
    using assms by(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   318
qed
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_NotL[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   322
  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
   323
  shows "pi1\<bullet>fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.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
   324
             fresh_fun (pi1\<bullet>(\<lambda>x'. Cut <c>.P (x').NotL <a>.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
   325
  and   "pi2\<bullet>fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.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
   326
             fresh_fun (pi2\<bullet>(\<lambda>x'. Cut <c>.P (x').NotL <a>.M x'))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   327
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   328
   apply(generate_fresh "name")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   329
   apply(auto simp: fresh_prod fresh_fun_simp_NotL)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   330
   apply(rule sym)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   331
   apply(rule trans)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   332
    apply(rule fresh_fun_simp_NotL)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   333
     apply(blast intro: fresh_perm_name)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   334
    apply(metis fresh_perm_name)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   335
   apply(simp add: at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   336
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   337
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   338
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   339
   apply (metis fresh_fun_simp_NotL fresh_prodD swap_simps(8) trm.perm(14) trm.perm(16))
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   340
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   343
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>x"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   344
  shows "fresh_fun (\<lambda>z'. Cut <c>.P(z').AndL1 (x).M z') = Cut <c>.P (z').AndL1 (x).M z'"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   345
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   346
  obtain n::name where "n\<sharp>(c,P,x,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   347
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   348
  then show "\<exists>a. a \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL1 x. M z', Cut <c>.P(a).AndL1 x. M a)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   349
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   350
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   351
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   352
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   353
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   354
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL1 x. M z')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   355
    using a by(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   356
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_AndL1[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   360
  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
   361
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).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
   362
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL1 (x).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
   363
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).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
   364
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL1 (x).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
   365
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
   366
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   367
apply(force simp add: fresh_prod fresh_fun_simp_AndL1 at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   368
  apply (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
   370
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   371
apply(force simp add: fresh_prod fresh_fun_simp_AndL1 calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   372
  by (meson exists_fresh'(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   375
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>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
   376
  shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z') = Cut <c>.P (z').AndL2 (x).M z'"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   377
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   378
  obtain n::name where "n\<sharp>(c,P,x,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   379
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   380
  then show "\<exists>a. a \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL2 x. M z', Cut <c>.P(a).AndL2 x. M a)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   381
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   382
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   383
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   384
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   385
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   386
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL2 x. M z')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   387
    using a by(fresh_guess)           
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   388
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_AndL2[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   392
  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
   393
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).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
   394
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL2 (x).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
   395
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).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
   396
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z'))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   397
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   398
   apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   399
    apply(force simp add: fresh_prod fresh_fun_simp_AndL2 at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   400
   apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   401
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   402
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   403
   apply(force simp add: fresh_prod fresh_fun_simp_AndL2 calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   404
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   407
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>N" "z'\<sharp>u" "z'\<sharp>x"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   408
  shows "fresh_fun (\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N z') = Cut <c>.P (z').OrL (x).M (u).N z'"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   409
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   410
  obtain n::name where "n\<sharp>(c,P,x,M,u,N)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   411
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   412
  then show "\<exists>a. a \<sharp> (\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N(z'), Cut <c>.P(a).OrL(x).M(u).N(a))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   413
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   414
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   415
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   416
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   417
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   418
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N(z'))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   419
    using a by(fresh_guess) 
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   420
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_OrL[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   424
  and   pi2::"coname prm"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   425
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P(z').OrL (x).M(u).N z')=
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   426
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').OrL(x).M (u).N z'))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   427
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P(z').OrL (x).M(u).N z')=
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   428
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N z'))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   429
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   430
   apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,u,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>N,pi1\<bullet>x,pi1\<bullet>u,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   431
    apply(force simp add: fresh_prod fresh_fun_simp_OrL at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   432
   apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   433
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   434
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,u,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>N,pi2\<bullet>x,pi2\<bullet>u,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   435
   apply(force simp add: fresh_prod fresh_fun_simp_OrL calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   436
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   439
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>N" "z'\<sharp>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
   440
  shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z') = Cut <c>.P (z').ImpL <a>.M (x).N z'"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   441
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   442
  obtain n::name where "n\<sharp>(c,P,x,M,N)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   443
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   444
  then show "\<exists>aa. aa \<sharp> (\<lambda>z'. Cut <c>.P(z').ImpL <a>.M(x).N z', Cut <c>.P(aa).ImpL <a>.M(x).N aa)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   445
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   446
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   447
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   448
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   449
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   450
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').ImpL <a>.M(x).N z')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   451
    using a by(fresh_guess) 
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   452
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_ImpL[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   456
  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
   457
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).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
   458
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).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
   459
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).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
   460
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z'))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   461
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   462
   apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>N,pi1\<bullet>x,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   463
    apply(force simp add: fresh_prod fresh_fun_simp_ImpL at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   464
   apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   465
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   466
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>N,pi2\<bullet>x,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   467
   apply(force simp add: fresh_prod fresh_fun_simp_ImpL calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   468
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   471
  assumes a: "a'\<sharp>P" "a'\<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
   472
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P) = Cut <a'>.(NotR (y).M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   473
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   474
  obtain n::coname where "n\<sharp>(x,P,y,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   475
    by (metis assms(1) assms(2) fresh_atm(4) fresh_prod)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   476
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P, Cut <a>.NotR(y).M(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   477
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   478
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   479
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   480
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   481
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to 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 fresh_fun_NotR[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   485
  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
   486
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (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
   487
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(NotR (y).M a') (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
   488
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (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
   489
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   490
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   491
   apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi1\<bullet>P,pi1\<bullet>M,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   492
    apply(force simp add: fresh_prod fresh_fun_simp_NotR calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   493
   apply (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   494
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   495
  apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   496
   apply(force simp: fresh_prod fresh_fun_simp_NotR at_prm_fresh[OF at_coname_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   497
  by (meson exists_fresh(2) fs_coname1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   500
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>N" "a'\<sharp>b" "a'\<sharp>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
   501
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P) = Cut <a'>.(AndR <b>.M <c>.N a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   502
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   503
  obtain n::coname where "n\<sharp>(x,P,b,M,c,N)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   504
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   505
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.AndR <b>.M <c>.N(a') (x).P, Cut <a>.AndR <b>.M <c>.N(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   506
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   507
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   508
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   509
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   510
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_AndR[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   514
  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
   515
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (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
   516
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (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
   517
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (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
   518
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   519
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   520
   apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,N,b,c,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>N,pi1\<bullet>b,pi1\<bullet>c,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   521
    apply(force simp add: fresh_prod fresh_fun_simp_AndR calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   522
   apply (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   523
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   524
  apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,N,b,c,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>N,pi2\<bullet>b,pi2\<bullet>c,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   525
   apply(force simp add: fresh_prod fresh_fun_simp_AndR at_prm_fresh[OF at_coname_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   526
  by (meson exists_fresh(2) fs_coname1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   529
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>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
   530
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P) = Cut <a'>.(OrR1 <b>.M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   531
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   532
  obtain n::coname where "n\<sharp>(x,P,b,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   533
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   534
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.OrR1 <b>.M(a') (x).P, Cut <a>.OrR1 <b>.M(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   535
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   536
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   537
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   538
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   539
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_OrR1[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   543
  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
   544
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (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
   545
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(OrR1 <b>.M  a') (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
   546
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (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
   547
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   548
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   549
   apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   550
    apply(force simp add: fresh_prod fresh_fun_simp_OrR1 calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   551
   apply (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   552
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   553
  apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   554
   apply(force simp add: fresh_prod fresh_fun_simp_OrR1 at_prm_fresh[OF at_coname_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   555
  by (meson exists_fresh(2) fs_coname1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   558
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>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
   559
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P) = Cut <a'>.(OrR2 <b>.M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   560
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   561
  obtain n::coname where "n\<sharp>(x,P,b,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   562
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   563
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.OrR2 <b>.M(a') (x).P, Cut <a>.OrR2 <b>.M(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   564
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   565
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   566
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   567
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   568
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_OrR2[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   572
  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
   573
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (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
   574
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(OrR2 <b>.M  a') (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
   575
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (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
   576
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   577
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   578
   apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   579
    apply(force simp add: fresh_prod fresh_fun_simp_OrR2 calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   580
   apply (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   581
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   582
  apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   583
   apply(force simp add: fresh_prod fresh_fun_simp_OrR2 at_prm_fresh[OF at_coname_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   584
  by (meson exists_fresh(2) fs_coname1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_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
   587
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>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
   588
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P) = Cut <a'>.(ImpR (y).<b>.M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   589
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   590
  obtain n::coname where "n\<sharp>(x,P,y,b,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   591
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   592
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P, Cut <a>.(ImpR (y).<b>.M a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   593
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   594
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   595
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   596
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   597
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   598
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_ImpR[eqvt_force]:
9be4ab2acc13 split Class.thy into parts to 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
  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
   601
  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
   602
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (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
   603
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M  a') (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
   604
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (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
   605
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   606
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   607
   apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   608
    apply(force simp add: fresh_prod fresh_fun_simp_ImpR calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   609
   apply (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   610
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   611
  apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   612
   apply(force simp add: fresh_prod fresh_fun_simp_ImpR at_prm_fresh[OF at_coname_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   613
  by (meson exists_fresh(2) fs_coname1)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
nominal_primrec (freshness_context: "(y::name,c::coname,P::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
   616
  substn :: "trm \<Rightarrow> name   \<Rightarrow> coname \<Rightarrow> trm \<Rightarrow> trm" ("_{_:=<_>._}" [100,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
   617
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
   618
  "(Ax x a){y:=<c>.P} = (if x=y then Cut <c>.P (y).Ax y a else 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
   619
| "\<lbrakk>a\<sharp>(c,P,N);x\<sharp>(y,P,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N){y:=<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
   620
  (if M=Ax y a then Cut <c>.P (x).(N{y:=<c>.P}) else Cut <a>.(M{y:=<c>.P}) (x).(N{y:=<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
   621
| "x\<sharp>(y,P) \<Longrightarrow> (NotR (x).M a){y:=<c>.P} = NotR (x).(M{y:=<c>.P}) 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
   622
| "a\<sharp>(c,P) \<Longrightarrow> (NotL <a>.M x){y:=<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
   623
  (if x=y then fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.(M{y:=<c>.P}) x') else NotL <a>.(M{y:=<c>.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
   624
| "\<lbrakk>a\<sharp>(c,P,N,d);b\<sharp>(c,P,M,d);b\<noteq>a\<rbrakk> \<Longrightarrow> 
9be4ab2acc13 split Class.thy into parts to 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
  (AndR <a>.M <b>.N d){y:=<c>.P} = AndR <a>.(M{y:=<c>.P}) <b>.(N{y:=<c>.P}) 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
   626
| "x\<sharp>(y,P,z) \<Longrightarrow> (AndL1 (x).M z){y:=<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
   627
  (if z=y then fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).(M{y:=<c>.P}) 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
   628
   else AndL1 (x).(M{y:=<c>.P}) 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
   629
| "x\<sharp>(y,P,z) \<Longrightarrow> (AndL2 (x).M z){y:=<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
   630
  (if z=y then fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).(M{y:=<c>.P}) 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
   631
   else AndL2 (x).(M{y:=<c>.P}) 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
   632
| "a\<sharp>(c,P,b) \<Longrightarrow> (OrR1 <a>.M b){y:=<c>.P} = OrR1 <a>.(M{y:=<c>.P}) 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
   633
| "a\<sharp>(c,P,b) \<Longrightarrow> (OrR2 <a>.M b){y:=<c>.P} = OrR2 <a>.(M{y:=<c>.P}) 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
   634
| "\<lbrakk>x\<sharp>(y,N,P,z);u\<sharp>(y,M,P,z);x\<noteq>u\<rbrakk> \<Longrightarrow> (OrL (x).M (u).N z){y:=<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
   635
  (if z=y then fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).(M{y:=<c>.P}) (u).(N{y:=<c>.P}) 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
   636
   else OrL (x).(M{y:=<c>.P}) (u).(N{y:=<c>.P}) 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
   637
| "\<lbrakk>a\<sharp>(b,c,P); x\<sharp>(y,P)\<rbrakk> \<Longrightarrow> (ImpR (x).<a>.M b){y:=<c>.P} = ImpR (x).<a>.(M{y:=<c>.P}) 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
   638
| "\<lbrakk>a\<sharp>(N,c,P);x\<sharp>(y,P,M,z)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N z){y:=<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
   639
  (if y=z then fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.(M{y:=<c>.P}) (x).(N{y:=<c>.P}) 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
   640
   else ImpL <a>.(M{y:=<c>.P}) (x).(N{y:=<c>.P}) z)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   641
apply(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   642
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   643
apply(force simp add: fresh_prod fresh_fun_simp_NotL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   644
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   645
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   646
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   647
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   648
apply(force simp add: fresh_prod fresh_fun_simp_AndL1 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   649
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   650
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   651
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   652
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   653
apply(force simp add: fresh_prod fresh_fun_simp_AndL2 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   654
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   655
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   656
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   657
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   658
apply(force simp add: fresh_prod fresh_fun_simp_OrL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   659
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   660
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   661
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   662
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   663
apply(force simp add: fresh_prod fresh_fun_simp_OrL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   664
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   665
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   666
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   667
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   668
apply(force simp add: fresh_prod fresh_fun_simp_ImpL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   669
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   670
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   671
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   672
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   673
apply(force simp add: fresh_prod fresh_fun_simp_ImpL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   674
apply (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
   676
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
   677
9be4ab2acc13 split Class.thy into parts to 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
nominal_primrec (freshness_context: "(d::name,z::coname,P::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
   679
  substc :: "trm \<Rightarrow> coname \<Rightarrow> name   \<Rightarrow> trm \<Rightarrow> trm" ("_{_:=(_)._}" [100,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
   680
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
   681
  "(Ax x a){d:=(z).P} = (if d=a then Cut <a>.(Ax x a) (z).P else 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
   682
| "\<lbrakk>a\<sharp>(d,P,N);x\<sharp>(z,P,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N){d:=(z).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
   683
  (if N=Ax x d then Cut <a>.(M{d:=(z).P}) (z).P else Cut <a>.(M{d:=(z).P}) (x).(N{d:=(z).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
   684
| "x\<sharp>(z,P) \<Longrightarrow> (NotR (x).M a){d:=(z).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
   685
  (if d=a then fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).(M{d:=(z).P}) a' (z).P) else NotR (x).(M{d:=(z).P}) 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
   686
| "a\<sharp>(d,P) \<Longrightarrow> (NotL <a>.M x){d:=(z).P} = NotL <a>.(M{d:=(z).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
   687
| "\<lbrakk>a\<sharp>(P,c,N,d);b\<sharp>(P,c,M,d);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c){d:=(z).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
   688
  (if d=c then fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.(M{d:=(z).P}) <b>.(N{d:=(z).P}) a') (z).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
   689
   else AndR <a>.(M{d:=(z).P}) <b>.(N{d:=(z).P}) 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
   690
| "x\<sharp>(y,z,P) \<Longrightarrow> (AndL1 (x).M y){d:=(z).P} = AndL1 (x).(M{d:=(z).P}) 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
   691
| "x\<sharp>(y,P,z) \<Longrightarrow> (AndL2 (x).M y){d:=(z).P} = AndL2 (x).(M{d:=(z).P}) 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
   692
| "a\<sharp>(d,P,b) \<Longrightarrow> (OrR1 <a>.M b){d:=(z).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
   693
  (if d=b then fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.(M{d:=(z).P}) a' (z).P) else OrR1 <a>.(M{d:=(z).P}) 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
   694
| "a\<sharp>(d,P,b) \<Longrightarrow> (OrR2 <a>.M b){d:=(z).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
   695
  (if d=b then fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.(M{d:=(z).P}) a' (z).P) else OrR2 <a>.(M{d:=(z).P}) 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
   696
| "\<lbrakk>x\<sharp>(N,z,P,u);y\<sharp>(M,z,P,u);x\<noteq>y\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N u){d:=(z).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
   697
  OrL (x).(M{d:=(z).P}) (y).(N{d:=(z).P}) 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
   698
| "\<lbrakk>a\<sharp>(b,d,P); x\<sharp>(z,P)\<rbrakk> \<Longrightarrow> (ImpR (x).<a>.M b){d:=(z).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
   699
  (if d=b then fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.(M{d:=(z).P}) a' (z).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
   700
   else ImpR (x).<a>.(M{d:=(z).P}) 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
   701
| "\<lbrakk>a\<sharp>(N,d,P);x\<sharp>(y,z,P,M)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y){d:=(z).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
   702
  ImpL <a>.(M{d:=(z).P}) (x).(N{d:=(z).P}) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   703
apply(finite_guess | simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   704
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   705
apply(force simp add: fresh_prod fresh_fun_simp_NotR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   706
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   707
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   708
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   709
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   710
apply(force simp add: fresh_prod fresh_fun_simp_AndR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   711
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   712
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   713
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   714
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   715
apply(force simp add: fresh_prod fresh_fun_simp_AndR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   716
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   717
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   718
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   719
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   720
apply(force simp add: fresh_prod fresh_fun_simp_OrR1 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   721
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   722
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   723
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   724
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   725
apply(force simp add: fresh_prod fresh_fun_simp_OrR2 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   726
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   727
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   728
apply(simp add: abs_fresh abs_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   729
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   730
apply(force simp add: fresh_prod fresh_fun_simp_ImpR abs_fresh fresh_atm abs_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   731
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   732
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   733
apply(simp add: abs_fresh abs_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   734
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   735
apply(force simp add: fresh_prod fresh_fun_simp_ImpR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   736
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   737
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   738
apply(simp add: abs_fresh | fresh_guess add: abs_fresh)+
36277
9be4ab2acc13 split Class.thy into parts to 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
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
   740
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   741
36277
9be4ab2acc13 split Class.thy into parts to 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
lemma csubst_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
   743
  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
   744
  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
   745
  shows "pi1\<bullet>(M{c:=(x).N}) = (pi1\<bullet>M){(pi1\<bullet>c):=(pi1\<bullet>x).(pi1\<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
   746
  and   "pi2\<bullet>(M{c:=(x).N}) = (pi2\<bullet>M){(pi2\<bullet>c):=(pi2\<bullet>x).(pi2\<bullet>N)}"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   747
by (nominal_induct M avoiding: c x N rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   748
   (auto simp: eq_bij fresh_bij eqvts; perm_simp)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma nsubst_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
   751
  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
   752
  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
   753
  shows "pi1\<bullet>(M{x:=<c>.N}) = (pi1\<bullet>M){(pi1\<bullet>x):=<(pi1\<bullet>c)>.(pi1\<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
   754
  and   "pi2\<bullet>(M{x:=<c>.N}) = (pi2\<bullet>M){(pi2\<bullet>x):=<(pi2\<bullet>c)>.(pi2\<bullet>N)}"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   755
by (nominal_induct M avoiding: c x N rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   756
   (auto simp: eq_bij fresh_bij eqvts; perm_simp)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst1:
9be4ab2acc13 split Class.thy into parts to 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
  shows "supp (M{y:=<c>.P}) \<subseteq> ((supp M) - {y}) \<union> (supp P)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   760
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   761
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   762
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   763
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   764
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   765
   show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   766
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   767
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   768
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   769
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   770
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   771
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   772
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   773
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   774
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   775
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   776
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   777
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   778
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   779
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   780
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   781
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   782
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   783
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   784
    by (auto simp: fs_name1 fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   785
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   786
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   787
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   788
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   789
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   790
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   791
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   792
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   793
text \<open>Identical to the previous proof\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst2:
9be4ab2acc13 split Class.thy into parts to 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
  shows "supp (M{y:=<c>.P}) \<subseteq> supp (M) \<union> ((supp P) - {c})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   796
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   797
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   798
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   799
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   800
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   801
   show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   802
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   803
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   804
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   805
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   806
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   807
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   808
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   809
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   810
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   811
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   812
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   813
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   814
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   815
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   816
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   817
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   818
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   819
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   820
    by (auto simp: fs_name1 fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   821
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   822
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   823
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   824
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   825
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   826
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   827
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst3:
9be4ab2acc13 split Class.thy into parts to 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
  shows "supp (M{c:=(x).P}) \<subseteq> ((supp M) - {c}) \<union> (supp P)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   831
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   832
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   833
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   834
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   835
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   836
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   837
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   838
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   839
  obtain x'::coname where x': "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   840
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   841
  with AndR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   842
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   843
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   844
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   845
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   846
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   847
  with OrR1 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   848
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   849
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   850
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   851
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   852
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   853
  with OrR2 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   854
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2; blast)  
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   855
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   856
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   857
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   858
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   859
  with ImpR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   860
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   861
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   862
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst4:
9be4ab2acc13 split Class.thy into parts to 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
  shows "supp (M{c:=(x).P}) \<subseteq> (supp M) \<union> ((supp P) - {x})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   866
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   867
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   868
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   869
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   870
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   871
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   872
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   873
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   874
  obtain x'::coname where x': "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   875
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   876
  with AndR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   877
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   878
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   879
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   880
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   881
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   882
  with OrR1 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   883
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   884
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   885
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   886
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   887
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   888
  with OrR2 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   889
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2; blast)  
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   890
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   891
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   892
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   893
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   894
  with ImpR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   895
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fs_name1 fresh_fun_simp_ImpR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   896
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst5:
9be4ab2acc13 split Class.thy into parts to 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
  shows "(supp M - {y}) \<subseteq> supp (M{y:=<c>.P})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   900
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   901
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   902
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   903
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   904
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   905
  show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   906
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   907
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   908
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   909
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   910
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   911
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   912
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   913
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   914
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   915
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   916
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   917
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   918
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   919
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   920
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   921
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   922
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   923
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   924
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   925
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   926
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   927
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   928
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   929
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   930
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   931
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   932
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   933
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   934
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   935
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   936
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   937
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   938
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   939
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   940
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   941
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst6:
9be4ab2acc13 split Class.thy into parts to 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
  shows "(supp M) \<subseteq> ((supp (M{y:=<c>.P}))::coname set)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   945
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   946
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   947
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   948
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   949
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   950
  show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   951
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   952
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   953
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   954
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   955
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   956
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   957
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   958
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   959
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   960
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   961
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   962
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   963
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   964
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   965
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   966
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   967
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   968
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   969
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   970
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   971
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   972
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   973
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   974
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   975
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   976
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   977
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   978
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   979
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   980
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   981
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   982
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   983
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   984
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   985
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   986
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst7:
9be4ab2acc13 split Class.thy into parts to 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
  shows "(supp M - {c}) \<subseteq>  supp (M{c:=(x).P})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   990
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   991
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   992
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   993
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   994
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   995
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   996
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   997
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   998
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   999
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1000
  obtain x'::coname where "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1001
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1002
  with AndR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1003
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1004
     apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1005
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1006
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1007
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1008
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1009
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1010
  with OrR1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1011
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1012
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1013
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1014
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1015
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1016
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1017
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1018
  with OrR2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1019
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1020
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1021
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1022
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1023
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1024
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1025
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1026
  with ImpR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1027
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1028
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1029
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1030
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1031
  
36277
9be4ab2acc13 split Class.thy into parts to 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
lemma supp_subst8:
9be4ab2acc13 split Class.thy into parts to 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
  shows "(supp M) \<subseteq> ((supp (M{c:=(x).P}))::name set)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1034
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1035
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1036
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1037
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1038
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1039
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1040
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1041
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1042
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1043
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1044
  obtain x'::coname where "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1045
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1046
  with AndR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1047
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1048
     apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1049
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1050
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1051
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1052
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1053
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1054
  with OrR1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1055
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1056
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1057
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1058
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1059
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1060
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1061
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1062
  with OrR2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1063
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1064
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1065
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1066
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1067
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1068
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1069
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1070
  with ImpR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1071
    by (force simp: fresh_prod abs_supp fs_name1 supp_atm trm.supp fresh_fun_simp_ImpR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1072
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1073
  
36277
9be4ab2acc13 split Class.thy into parts to 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
lemmas subst_supp = supp_subst1 supp_subst2 supp_subst3 supp_subst4
9be4ab2acc13 split Class.thy into parts to 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
                    supp_subst5 supp_subst6 supp_subst7 supp_subst8
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1077
  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
  1078
  and   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
  1079
  shows "x\<sharp>P \<Longrightarrow> x\<sharp>M{x:=<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
  1080
  and   "b\<sharp>P \<Longrightarrow> b\<sharp>M{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
  1081
  and   "x\<sharp>(M,P) \<Longrightarrow> x\<sharp>M{y:=<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
  1082
  and   "x\<sharp>M \<Longrightarrow> x\<sharp>M{c:=(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
  1083
  and   "x\<sharp>(M,P) \<Longrightarrow> x\<sharp>M{c:=(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
  1084
  and   "b\<sharp>(M,P) \<Longrightarrow> b\<sharp>M{c:=(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
  1085
  and   "b\<sharp>M \<Longrightarrow> b\<sharp>M{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
  1086
  and   "b\<sharp>(M,P) \<Longrightarrow> b\<sharp>M{y:=<c>.P}"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1087
  using subst_supp
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1088
  by(fastforce simp add: fresh_def supp_prod)+
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1091
  shows "x\<sharp>M \<Longrightarrow> M{x:=<c>.P} = 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
  1092
  and   "c\<sharp>M \<Longrightarrow> M{c:=(x).P} = 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
  1093
apply(nominal_induct M avoiding: x c P rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1094
apply(auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1096
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_rename1:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "c\<sharp>(M,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
  1099
  shows "M{a:=(x).N} = ([(c,a)]\<bullet>M){c:=(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
  1100
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
  1101
proof(nominal_induct M avoiding: c a 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
  1102
  case (AndR c1 M c2 M' c3)
9be4ab2acc13 split Class.thy into parts to 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
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1104
    apply(auto simp: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1105
     apply (metis (no_types, lifting))+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1106
    done
36277
9be4ab2acc13 split Class.thy into parts to 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
next 
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  then show ?case
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1110
    by (auto simp: calc_atm alpha fresh_atm abs_fresh fresh_prod fresh_left)
56073
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 53015
diff changeset
  1111
       metis
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut d M y 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
  1114
  then show ?case
56073
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 53015
diff changeset
  1115
    by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 53015
diff changeset
  1116
      (metis crename.simps(1) crename_id crename_rename)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1117
qed (auto simp: calc_atm alpha fresh_atm abs_fresh fresh_prod fresh_left trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_rename2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "y\<sharp>(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
  1121
  shows "M{a:=(x).N} = M{a:=(y).([(y,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
  1122
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
  1123
proof(nominal_induct M avoiding: a 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
  1124
  case (NotR y M d)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1125
  obtain a::coname where "a\<sharp>(N,M{d:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1126
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1127
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1128
    apply(auto simp: calc_atm alpha fresh_atm fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1129
    by (metis (no_types, opaque_lifting) alpha(1) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR c1 M c2 M' c3)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1132
  obtain a'::coname where "a'\<sharp>(N,M{c3:=(y).([(y,x)]\<bullet>N)},M'{c3:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,c1,c2,c3)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1133
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1134
  with AndR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1135
    apply(auto simp: calc_atm alpha fresh_atm fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1136
    by (metis (no_types, opaque_lifting) alpha(1) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR1 d M e)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1139
  obtain a'::coname where "a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1140
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1141
  with OrR1 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1142
    by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_OrR1)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR2 d M e)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1145
  obtain a'::coname where "a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1146
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1147
  with OrR2 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1148
    by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_OrR2)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpR y d M e)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1151
  obtain a'::coname where "a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1152
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1153
  with ImpR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1154
    by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_ImpR) 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1155
qed (auto simp: calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_rename3:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "y\<sharp>(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
  1159
  shows "M{x:=<a>.N} = ([(y,x)]\<bullet>M){y:=<a>.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
  1160
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
  1161
proof(nominal_induct M avoiding: a 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
  1162
  case (OrL x1 M x2 M' 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
  1163
  then show ?case
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1164
    apply(auto simp add: calc_atm fresh_atm abs_fresh fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1165
    by (metis (mono_tags))+
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL d M v M' 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
  1168
  then show ?case
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1169
    apply(auto simp add: calc_atm fresh_atm abs_fresh fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1170
    by (metis (mono_tags))+
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut d M y 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
  1173
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1174
    apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1175
    by (metis nrename.simps(1) nrename_id nrename_rename)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1176
qed (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_left abs_supp fin_supp fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_rename4:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "c\<sharp>(N,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
  1180
  shows "M{x:=<a>.N} = M{x:=<c>.([(c,a)]\<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
  1181
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
  1182
proof(nominal_induct M avoiding: x c a 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
  1183
  case (Ax z 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
  1184
  then show ?case 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1185
    by (auto simp: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  1188
  then show ?case 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1189
    by (auto simp: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL d 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
  1192
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1193
    apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<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
  1195
    apply(erule exE, simp add: 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
  1196
    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
  1197
    apply(simp add: fresh_fun_simp_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
  1198
    apply(simp add: trm.inject alpha perm_swap 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
  1199
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  1201
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL x1 M x2 M' 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
  1203
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1204
    apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(subgoal_tac 
9be4ab2acc13 split Class.thy into parts to 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
       "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},M'{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,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
  1207
    apply(erule exE, simp add: 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
  1208
    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
  1209
    apply(simp add: fresh_fun_simp_OrL)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1210
    apply (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh perm_swap fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  1213
next
9be4ab2acc13 split Class.thy into parts to 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
  case OrR1
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1215
  then show ?case by (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case OrR2
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1218
  then show ?case by (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 u M 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
  1221
  then show ?case 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1222
    apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,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
  1224
    apply(erule exE, simp add: 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
  1225
    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
  1226
    apply(simp add: fresh_fun_simp_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
  1227
    apply(simp add: trm.inject alpha perm_swap 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
  1228
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  1230
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL2 u M 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
  1232
  then show ?case 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1233
    apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,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
  1235
    apply(erule exE, simp add: 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
  1236
    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
  1237
    apply(simp add: fresh_fun_simp_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
  1238
    apply(simp add: trm.inject alpha perm_swap 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
  1239
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  1241
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR c1 M c2 M' c3)
9be4ab2acc13 split Class.thy into parts to 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
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1244
    by (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
next 
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  1247
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1248
    by (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL d M y M' 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
  1251
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1252
    apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{u:=<c>.([(c,a)]\<bullet>N)},M'{u:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,y,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
  1254
    apply(erule exE, simp add: 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
  1255
    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
  1256
    apply(simp add: fresh_fun_simp_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
  1257
    apply(simp add: trm.inject alpha perm_swap 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
  1258
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  1260
next
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut d M y 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
  1262
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1263
    by (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1265
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_rename5:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "c'\<sharp>(c,N)" "x'\<sharp>(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
  1268
  shows "M{x:=<c>.N} = ([(x',x)]\<bullet>M){x':=<c'>.([(c',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
  1269
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
  1270
  have "M{x:=<c>.N} = ([(x',x)]\<bullet>M){x':=<c>.N}" using a by (simp add: substn_rename3)
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(x',x)]\<bullet>M){x':=<c'>.([(c',c)]\<bullet>N)}" using a by (simp add: substn_rename4)
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  1274
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_rename6:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "c'\<sharp>(c,M)" "x'\<sharp>(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
  1277
  shows "M{c:=(x).N} = ([(c',c)]\<bullet>M){c':=(x').([(x',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
  1278
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
  1279
  have "M{c:=(x).N} = ([(c',c)]\<bullet>M){c':=(x).N}" using a by (simp add: substc_rename1)
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(c',c)]\<bullet>M){c':=(x').([(x',x)]\<bullet>N)}" using a by (simp add: substc_rename2)
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  1283
9be4ab2acc13 split Class.thy into parts to 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
lemmas subst_rename = substc_rename1 substc_rename2 substn_rename3 substn_rename4 subst_rename5 subst_rename6
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma better_Cut_substn[simp]:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "a\<sharp>[c].P" "x\<sharp>(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
  1288
  shows "(Cut <a>.M (x).N){y:=<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
  1289
  (if M=Ax y a then Cut <c>.P (x).(N{y:=<c>.P}) else Cut <a>.(M{y:=<c>.P}) (x).(N{y:=<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
  1290
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
  1291
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,c,P,x,y)" by (rule exists_fresh(1), rule fin_supp, 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
  1292
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,c,P,a)" by (rule exists_fresh(2), rule fin_supp, 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
  1293
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1294
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  have eq2: "(M=Ax y a) = (([(a',a)]\<bullet>M)=Ax y a')"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1296
    apply(auto simp: calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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 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
  1298
    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
  1299
    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
  1300
  have "(Cut <a>.M (x).N){y:=<c>.P} = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)){y:=<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
  1301
    using eq1 by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = (if ([(a',a)]\<bullet>M)=Ax y a' then Cut <c>.P (x').(([(x',x)]\<bullet>N){y:=<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
  1303
                              else Cut <a'>.(([(a',a)]\<bullet>M){y:=<c>.P}) (x').(([(x',x)]\<bullet>N){y:=<c>.P}))" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1304
    using fs1 fs2 by (auto simp: fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> =(if M=Ax y a then Cut <c>.P (x).(N{y:=<c>.P}) else Cut <a>.(M{y:=<c>.P}) (x).(N{y:=<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
  1306
    using fs1 fs2 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
  1307
    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
  1308
    apply(simp only: eq2[symmetric])
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1309
    apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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_all add: alpha fresh_atm fresh_prod 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
  1311
    apply(simp_all add: eqvts perm_fresh_fresh 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
  1312
    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
  1313
    apply(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(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
  1315
    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
  1316
    apply(simp add: perm_fresh_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
  1317
    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
  1318
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  1320
    
9be4ab2acc13 split Class.thy into parts to 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
lemma better_Cut_substc[simp]:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "a\<sharp>(c,P)" "x\<sharp>[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
  1323
  shows "(Cut <a>.M (x).N){c:=(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
  1324
  (if N=Ax x c then Cut <a>.(M{c:=(y).P}) (y).P else Cut <a>.(M{c:=(y).P}) (x).(N{c:=(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
  1325
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
  1326
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,c,P,x,y)" by (rule exists_fresh(1), rule fin_supp, 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
  1327
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,c,P,a)" by (rule exists_fresh(2), rule fin_supp, 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
  1328
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1329
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  have eq2: "(N=Ax x c) = (([(x',x)]\<bullet>N)=Ax x' c)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1331
    apply(auto simp: calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  1333
    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
  1334
    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
  1335
  have "(Cut <a>.M (x).N){c:=(y).P} = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)){c:=(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
  1336
    using eq1 by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = (if ([(x',x)]\<bullet>N)=Ax x' c then  Cut <a'>.(([(a',a)]\<bullet>M){c:=(y).P}) (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
  1338
                              else Cut <a'>.(([(a',a)]\<bullet>M){c:=(y).P}) (x').(([(x',x)]\<bullet>N){c:=(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
  1339
    using fs1 fs2  by (simp add: fresh_prod fresh_left calc_atm fresh_atm 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
  1340
  also have "\<dots> =(if N=Ax x c then Cut <a>.(M{c:=(y).P}) (y).P else Cut <a>.(M{c:=(y).P}) (x).(N{c:=(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
  1341
    using fs1 fs2 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
  1342
    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
  1343
    apply(simp only: eq2[symmetric])
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1344
    apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(simp_all add: alpha fresh_atm fresh_prod 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
  1346
    apply(simp_all add: eqvts perm_fresh_fresh 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
  1347
    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
  1348
    apply(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(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
  1350
    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
  1351
    apply(simp add: perm_fresh_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
  1352
    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
  1353
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  1355
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1357
  assumes a: "a\<sharp>[c].P" "y\<sharp>(N,x)" "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
  1358
  shows "(Cut <a>.M (x).N){y:=<c>.P} = Cut <a>.(M{y:=<c>.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
  1359
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
  1360
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
  1361
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
  1362
apply(subgoal_tac "Cut <a>.M (x).N = Cut <a>.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
  1363
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  1365
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
  1366
apply(simp add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1367
apply(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1369
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
  1370
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
  1371
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
  1372
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1374
  assumes a: "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
  1375
  shows "(NotR (x).M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).M a' (z).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
  1376
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
  1377
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
  1378
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
  1379
apply(subgoal_tac "NotR (x).M d = NotR (c).([(c,x)]\<bullet>M) d")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1380
apply(auto simp: fresh_left calc_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1382
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1383
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1385
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
  1386
apply(perm_simp add: trm.inject alpha fresh_prod fresh_atm fresh_left, 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
  1387
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
  1388
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1390
  assumes a: "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
  1391
  shows "(NotL <a>.M y){y:=<c>.P} = fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.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
  1392
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
  1393
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
  1394
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
  1395
apply(subgoal_tac "NotL <a>.M y = NotL <ca>.([(ca,a)]\<bullet>M) y")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1396
apply(auto simp: fresh_left calc_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1398
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1399
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1401
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
  1402
apply(perm_simp add: trm.inject alpha fresh_prod fresh_atm fresh_left, 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
  1403
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
  1404
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1406
  assumes a: "y\<sharp>[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
  1407
  shows "(AndL1 (x).M y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).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
  1408
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
  1409
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
  1410
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
  1411
apply(subgoal_tac "AndL1 (x).M y = AndL1 (ca).([(ca,x)]\<bullet>M) y")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1412
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1414
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1415
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1417
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
  1418
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
  1419
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
  1420
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1421
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1423
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
  1424
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
  1425
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
  1426
apply(perm_simp add: trm.inject alpha 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
  1427
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
  1428
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
  1429
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1431
  assumes a: "y\<sharp>[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
  1432
  shows "(AndL2 (x).M y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).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
  1433
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
  1434
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
  1435
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
  1436
apply(subgoal_tac "AndL2 (x).M y = AndL2 (ca).([(ca,x)]\<bullet>M) y")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1437
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1439
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1440
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1442
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
  1443
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
  1444
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
  1445
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1446
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1448
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
  1449
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
  1450
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
  1451
apply(perm_simp add: trm.inject alpha 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
  1452
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
  1453
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
  1454
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1456
  assumes a: "c\<sharp>([a].M,[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
  1457
  shows "(AndR <a>.M <b>.N c){c:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.M <b>.N a') (z).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
  1458
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
  1459
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
  1460
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
  1461
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
  1462
apply(subgoal_tac "AndR <a>.M <b>.N c = AndR <ca>.([(ca,a)]\<bullet>M) <caa>.([(caa,b)]\<bullet>N) c")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1463
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1465
apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1466
apply(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1467
apply(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1468
apply(auto simp: fresh_prod fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1471
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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 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
  1473
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
  1474
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
  1475
apply(rule forget)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1476
apply(auto simp: fresh_left calc_atm abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule forget)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1478
apply(auto simp: fresh_left calc_atm abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(perm_simp add: trm.inject alpha 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
  1480
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
  1481
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
  1482
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1484
  assumes a: "x\<sharp>([y].M,[z].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
  1485
  shows "(OrL (y).M (z).N x){x:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (y).M (z).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
  1486
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
  1487
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
  1488
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
  1489
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
  1490
apply(subgoal_tac "OrL (y).M (z).N x = OrL (ca).([(ca,y)]\<bullet>M) (caa).([(caa,z)]\<bullet>N) x")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1491
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1493
apply(rule substn.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1494
apply(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1495
apply(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1496
apply(auto simp: fresh_prod fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1499
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1501
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
  1502
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
  1503
apply(rule forget)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1504
apply(auto simp: fresh_left calc_atm abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule forget)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1506
apply(auto simp: fresh_left calc_atm abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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 add: trm.inject alpha 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
  1508
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
  1509
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
  1510
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1512
  assumes a: "d\<sharp>[a].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
  1513
  shows "(OrR1 <a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.M a' (z).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
  1514
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
  1515
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
  1516
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
  1517
apply(subgoal_tac "OrR1 <a>.M d = OrR1 <c>.([(c,a)]\<bullet>M) d")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1518
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1520
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1522
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
  1523
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
  1524
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
  1525
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1526
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1528
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
  1529
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
  1530
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
  1531
apply(perm_simp add: trm.inject alpha 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
  1532
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
  1533
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
  1534
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1536
  assumes a: "d\<sharp>[a].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
  1537
  shows "(OrR2 <a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.M a' (z).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
  1538
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
  1539
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
  1540
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
  1541
apply(subgoal_tac "OrR2 <a>.M d = OrR2 <c>.([(c,a)]\<bullet>M) d")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1542
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1544
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1546
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
  1547
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
  1548
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
  1549
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1550
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1552
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
  1553
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
  1554
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
  1555
apply(perm_simp add: trm.inject alpha 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
  1556
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
  1557
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
  1558
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  1560
  assumes a: "d\<sharp>[a].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
  1561
  shows "(ImpR (x).<a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.M a' (z).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
  1562
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
  1563
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
  1564
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
  1565
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
  1566
apply(subgoal_tac "ImpR (x).<a>.M d = ImpR (ca).<c>.([(c,a)]\<bullet>[(ca,x)]\<bullet>M) d")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1567
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1569
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1571
apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm abs_fresh 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
  1572
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
  1573
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
  1574
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1575
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1577
apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm fresh_left calc_atm 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
  1578
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
  1579
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
  1580
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
  1581
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_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
  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 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
  1585
  assumes a: "y\<sharp>(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
  1586
  shows "(ImpL <a>.M (x).N y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).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
  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 -
9be4ab2acc13 split Class.thy into parts to 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(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
  1590
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
  1591
apply(subgoal_tac "ImpL <a>.M (x).N y = ImpL <ca>.([(ca,a)]\<bullet>M) (caa).([(caa,x)]\<bullet>N) y")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1592
apply(auto simp: fresh_left calc_atm forget abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac f="fresh_fun" in arg_cong)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1594
apply(simp add:  fun_eq_iff)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1596
apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm abs_fresh 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
  1597
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
  1598
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
  1599
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
  1600
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
  1601
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_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
  1602
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
  1603
9be4ab2acc13 split Class.thy into parts to 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
lemma freshn_after_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
  1605
  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
  1606
  assumes a: "x\<sharp>M{c:=(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
  1607
  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
  1608
using a supp_subst8
9be4ab2acc13 split Class.thy into parts to 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
apply(simp add: fresh_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
  1610
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
  1611
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
  1612
9be4ab2acc13 split Class.thy into parts to 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
lemma freshn_after_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
  1614
  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
  1615
  assumes a: "x\<sharp>M{y:=<c>.P}" "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
  1616
  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
  1617
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
  1618
using a supp_subst5
9be4ab2acc13 split Class.thy into parts to 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(simp add: fresh_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
  1620
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
  1621
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
  1622
9be4ab2acc13 split Class.thy into parts to 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
lemma freshc_after_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
  1624
  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
  1625
  assumes a: "a\<sharp>M{c:=(y).P}" "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
  1626
  shows "a\<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
  1627
using a supp_subst7
9be4ab2acc13 split Class.thy into parts to 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(simp add: fresh_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
  1629
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
  1630
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
  1631
9be4ab2acc13 split Class.thy into parts to 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
lemma freshc_after_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
  1633
  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
  1634
  assumes a: "a\<sharp>M{y:=<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
  1635
  shows "a\<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
  1636
using a supp_subst6
9be4ab2acc13 split Class.thy into parts to 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(simp add: fresh_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
  1638
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
  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 substn_crename_comm:
9be4ab2acc13 split Class.thy into parts to 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: "c\<noteq>a" "c\<noteq>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
  1643
  shows "M{x:=<c>.P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{x:=<c>.(P[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
  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 M avoiding: x c P a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1646
apply(auto simp: subst_fresh rename_fresh trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,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
  1648
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
  1649
apply(subgoal_tac "Cut <c>.P (x).Ax x a = 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
  1650
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  1652
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
  1653
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
  1654
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  1656
apply(simp add: alpha trm.inject 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
  1657
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
  1658
apply(simp add: alpha trm.inject 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
  1659
apply(rule exists_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
  1660
apply(rule 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
  1661
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
  1662
apply(rule better_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
  1663
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
  1664
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  1666
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
  1667
apply(rule better_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
  1668
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
  1669
apply(simp add: rename_fresh fresh_atm)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1670
apply(auto simp: fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1672
apply(rule better_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
  1673
apply(simp add: fresh_atm)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1674
apply(auto simp: fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(drule crename_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
  1676
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
  1677
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
  1678
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[a\<turnstile>c>b],x,trm[a\<turnstile>c>b]{x:=<c>.P[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
  1680
apply(erule exE, simp add: 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
  1681
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
  1682
apply(simp add: fresh_fun_simp_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
  1683
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
  1684
apply(rule better_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
  1685
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
  1686
apply(simp add: rename_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
  1687
apply(rule exists_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
  1688
apply(rule 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
  1689
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[a\<turnstile>c>b],name1,trm[a\<turnstile>c>b]{x:=<c>.P[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
  1690
apply(erule exE, simp add: 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
  1691
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
  1692
apply(simp add: fresh_fun_simp_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
  1693
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
  1694
apply(rule better_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
  1695
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
  1696
apply(simp add: rename_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
  1697
apply(rule exists_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
  1698
apply(rule 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
  1699
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[a\<turnstile>c>b],name1,trm[a\<turnstile>c>b]{x:=<c>.P[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
  1700
apply(erule exE, simp add: 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
  1701
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
  1702
apply(simp add: fresh_fun_simp_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
  1703
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
  1704
apply(rule better_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
  1705
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
  1706
apply(simp add: rename_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
  1707
apply(rule exists_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
  1708
apply(rule 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
  1709
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[a\<turnstile>c>b],name1,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
  1710
                                  trm1[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]},trm2[a\<turnstile>c>b]{x:=<c>.P[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
  1711
apply(erule exE, simp add: 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
  1712
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
  1713
apply(simp add: fresh_fun_simp_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
  1714
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
  1715
apply(rule better_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
  1716
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
  1717
apply(simp add: rename_fresh subst_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
  1718
apply(rule exists_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
  1719
apply(rule 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
  1720
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},P,P[a\<turnstile>c>b],name1,
9be4ab2acc13 split Class.thy into parts to 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
                                  trm1[a\<turnstile>c>b]{name2:=<c>.P[a\<turnstile>c>b]},trm2[a\<turnstile>c>b]{name2:=<c>.P[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
  1722
apply(erule exE, simp add: 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
  1723
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
  1724
apply(simp add: fresh_fun_simp_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
  1725
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
  1726
apply(rule better_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
  1727
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
  1728
apply(simp add: rename_fresh subst_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
  1729
apply(rule exists_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
  1730
apply(rule 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
  1731
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
  1732
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_crename_comm:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "c\<noteq>a" "c\<noteq>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
  1735
  shows "M{c:=(x).P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{c:=(x).(P[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
  1736
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
  1737
apply(nominal_induct M avoiding: x c P a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1738
apply(auto simp: subst_fresh rename_fresh trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1740
apply(rule better_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
  1741
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
  1742
apply(simp add: rename_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
  1743
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
  1744
apply(rule better_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
  1745
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
  1746
apply(simp add: rename_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(drule crename_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
  1748
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
  1749
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
  1750
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>c'::coname. c'\<sharp>(a,b,trm{coname:=(x).P},P,P[a\<turnstile>c>b],x,trm[a\<turnstile>c>b]{coname:=(x).P[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
  1752
apply(erule exE, simp add: 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
  1753
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
  1754
apply(simp add: fresh_fun_simp_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
  1755
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
  1756
apply(rule better_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
  1757
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
  1758
apply(simp add: rename_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
  1759
apply(rule exists_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
  1760
apply(rule 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
  1761
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,coname2,a,b,trm1{coname3:=(x).P},trm2{coname3:=(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
  1762
                   P,P[a\<turnstile>c>b],x,trm1[a\<turnstile>c>b]{coname3:=(x).P[a\<turnstile>c>b]},trm2[a\<turnstile>c>b]{coname3:=(x).P[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
  1763
apply(erule exE, simp add: 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
  1764
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
  1765
apply(simp add: fresh_fun_simp_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
  1766
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
  1767
apply(rule better_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
  1768
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
  1769
apply(simp add: rename_fresh subst_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
  1770
apply(rule exists_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
  1771
apply(rule 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
  1772
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],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
  1773
                         trm[a\<turnstile>c>b]{coname2:=(x).P[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
  1774
apply(erule exE, simp add: 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
  1775
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
  1776
apply(simp add: fresh_fun_simp_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
  1777
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
  1778
apply(rule better_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
  1779
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
  1780
apply(simp add: rename_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
  1781
apply(rule exists_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
  1782
apply(rule 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
  1783
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],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
  1784
                         trm[a\<turnstile>c>b]{coname2:=(x).P[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
  1785
apply(erule exE, simp add: 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
  1786
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
  1787
apply(simp add: fresh_fun_simp_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
  1788
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
  1789
apply(rule better_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
  1790
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
  1791
apply(simp add: rename_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
  1792
apply(rule exists_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
  1793
apply(rule 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
  1794
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],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
  1795
                         trm[a\<turnstile>c>b]{coname2:=(x).P[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
  1796
apply(erule exE, simp add: 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
  1797
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
  1798
apply(simp add: fresh_fun_simp_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
  1799
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
  1800
apply(rule better_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
  1801
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
  1802
apply(simp add: rename_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
  1803
apply(rule exists_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
  1804
apply(rule 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
  1805
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
  1806
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_nrename_comm:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "x\<noteq>y" "x\<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
  1809
  shows "M{x:=<c>.P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{x:=<c>.(P[y\<turnstile>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
  1810
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
  1811
apply(nominal_induct M avoiding: x c P y z rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1812
apply(auto simp: subst_fresh rename_fresh trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1814
apply(rule better_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
  1815
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
  1816
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
  1817
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
  1818
apply(rule better_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_atm)
9be4ab2acc13 split Class.thy into parts to 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)
9be4ab2acc13 split Class.thy into parts to 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(drule nrename_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
  1822
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
  1823
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
  1824
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,z,trm{x:=<c>.P},P,P[y\<turnstile>n>z],x,trm[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>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
  1826
apply(erule exE, simp add: 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
  1827
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
  1828
apply(simp add: fresh_fun_simp_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
  1829
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
  1830
apply(rule better_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
  1831
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
  1832
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
  1833
apply(rule exists_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
  1834
apply(rule 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
  1835
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[y\<turnstile>n>z],name1,trm[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]},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
  1836
apply(erule exE, simp add: 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
  1837
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
  1838
apply(simp add: fresh_fun_simp_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
  1839
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
  1840
apply(rule better_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
  1841
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
  1842
apply(simp add: rename_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
  1843
apply(rule exists_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
  1844
apply(rule 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
  1845
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,z,trm{x:=<c>.P},P,P[y\<turnstile>n>z],name1,trm[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>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
  1846
apply(erule exE, simp add: 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
  1847
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
  1848
apply(simp add: fresh_fun_simp_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
  1849
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
  1850
apply(rule better_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
  1851
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
  1852
apply(simp add: rename_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
  1853
apply(rule exists_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
  1854
apply(rule 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
  1855
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[y\<turnstile>n>z],name1,name2,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
  1856
                                  trm1[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]},trm2[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>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
  1857
apply(erule exE, simp add: 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
  1858
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
  1859
apply(simp add: fresh_fun_simp_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
  1860
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
  1861
apply(rule better_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
  1862
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
  1863
apply(simp add: rename_fresh subst_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
  1864
apply(rule exists_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
  1865
apply(rule 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
  1866
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},P,P[y\<turnstile>n>z],y,z,name1,
9be4ab2acc13 split Class.thy into parts to 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
                                  trm1[y\<turnstile>n>z]{name2:=<c>.P[y\<turnstile>n>z]},trm2[y\<turnstile>n>z]{name2:=<c>.P[y\<turnstile>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
  1868
apply(erule exE, simp add: 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
  1869
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
  1870
apply(simp add: fresh_fun_simp_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
  1871
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
  1872
apply(rule better_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
  1873
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
  1874
apply(simp add: rename_fresh subst_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
  1875
apply(rule exists_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
  1876
apply(rule 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
  1877
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
  1878
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_nrename_comm:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "x\<noteq>y" "x\<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
  1881
  shows "M{c:=(x).P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{c:=(x).(P[y\<turnstile>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
  1882
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
  1883
apply(nominal_induct M avoiding: x c P y z rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1884
apply(auto simp: subst_fresh rename_fresh trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  1886
apply(rule better_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
  1887
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
  1888
apply(simp add: rename_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
  1889
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
  1890
apply(rule better_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
  1891
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
  1892
apply(simp add: rename_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
  1893
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
  1894
apply(rule better_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
  1895
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
  1896
apply(simp add: rename_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
  1897
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
  1898
apply(rule better_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
  1899
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
  1900
apply(simp add: rename_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
  1901
apply(drule nrename_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
  1902
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
  1903
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
  1904
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  1906
apply(rule better_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
  1907
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
  1908
apply(simp add: rename_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
  1909
apply(drule nrename_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
  1910
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
  1911
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
  1912
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>c'::coname. c'\<sharp>(y,z,trm{coname:=(x).P},P,P[y\<turnstile>n>z],x,trm[y\<turnstile>n>z]{coname:=(x).P[y\<turnstile>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
  1914
apply(erule exE, simp add: 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
  1915
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
  1916
apply(simp add: fresh_fun_simp_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
  1917
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
  1918
apply(rule better_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
  1919
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
  1920
apply(simp add: rename_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
  1921
apply(rule exists_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
  1922
apply(rule 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
  1923
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,coname2,y,z,trm1{coname3:=(x).P},trm2{coname3:=(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
  1924
                   P,P[y\<turnstile>n>z],x,trm1[y\<turnstile>n>z]{coname3:=(x).P[y\<turnstile>n>z]},trm2[y\<turnstile>n>z]{coname3:=(x).P[y\<turnstile>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
  1925
apply(erule exE, simp add: 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
  1926
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
  1927
apply(simp add: fresh_fun_simp_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
  1928
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
  1929
apply(rule better_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
  1930
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
  1931
apply(simp add: rename_fresh subst_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
  1932
apply(rule exists_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
  1933
apply(rule 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
  1934
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],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
  1935
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>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
  1936
apply(erule exE, simp add: 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
  1937
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
  1938
apply(simp add: fresh_fun_simp_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
  1939
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
  1940
apply(rule better_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
  1941
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
  1942
apply(simp add: rename_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
  1943
apply(rule exists_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
  1944
apply(rule 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
  1945
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],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
  1946
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>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
  1947
apply(erule exE, simp add: 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
  1948
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
  1949
apply(simp add: fresh_fun_simp_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
  1950
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
  1951
apply(rule better_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
  1952
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
  1953
apply(simp add: rename_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
  1954
apply(rule exists_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
  1955
apply(rule 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
  1956
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],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
  1957
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>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
  1958
apply(erule exE, simp add: 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
  1959
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
  1960
apply(simp add: fresh_fun_simp_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
  1961
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
  1962
apply(rule better_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
  1963
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
  1964
apply(simp add: rename_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
  1965
apply(rule exists_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
  1966
apply(rule 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
  1967
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
  1968
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_crename_comm':
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "a\<noteq>c" "a\<sharp>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
  1971
  shows "M{x:=<c>.P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{x:=<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
  1972
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
  1973
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
  1974
  assume a1: "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
  1975
  assume a2: "a\<sharp>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
  1976
  obtain c'::"coname" where fs2: "c'\<sharp>(c,P,a,b)" by (rule exists_fresh(2), rule fin_supp, 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
  1977
  have eq: "M{x:=<c>.P} = M{x:=<c'>.([(c',c)]\<bullet>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
  1978
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  1983
   have eq': "M[a\<turnstile>c>b]{x:=<c>.P} = M[a\<turnstile>c>b]{x:=<c'>.([(c',c)]\<bullet>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
  1984
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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 subst_rename)
9be4ab2acc13 split Class.thy into parts to 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)
9be4ab2acc13 split Class.thy into parts to 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
    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
  1989
  have eq2: "([(c',c)]\<bullet>P)[a\<turnstile>c>b] = ([(c',c)]\<bullet>P)" using fs2 a2 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
  1990
    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
  1991
    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
  1992
    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
  1993
    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
  1994
  have "M{x:=<c>.P}[a\<turnstile>c>b] = M{x:=<c'>.([(c',c)]\<bullet>P)}[a\<turnstile>c>b]" using eq by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = M[a\<turnstile>c>b]{x:=<c'>.(([(c',c)]\<bullet>P)[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
  1996
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule substn_crename_comm)
9be4ab2acc13 split Class.thy into parts to 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(simp_all 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
  2000
    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
  2001
  also have "\<dots> = M[a\<turnstile>c>b]{x:=<c'>.(([(c',c)]\<bullet>P))}" using eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = M[a\<turnstile>c>b]{x:=<c>.P}" using eq' by simp 
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  2005
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_crename_comm':
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "c\<noteq>a" "c\<noteq>b" "a\<sharp>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
  2008
  shows "M{c:=(x).P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{c:=(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
  2009
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
  2010
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
  2011
  assume a1: "c\<noteq>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
  2012
  assume a1': "c\<noteq>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
  2013
  assume a2: "a\<sharp>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
  2014
  obtain c'::"coname" where fs2: "c'\<sharp>(c,M,a,b)" by (rule exists_fresh(2), rule fin_supp, 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
  2015
  have eq: "M{c:=(x).P} = ([(c',c)]\<bullet>M){c':=(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
  2016
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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 subst_rename)
9be4ab2acc13 split Class.thy into parts to 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)
9be4ab2acc13 split Class.thy into parts to 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
    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
  2021
   have eq': "([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P} = M[a\<turnstile>c>b]{c:=(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
  2022
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule subst_rename[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
  2025
    apply(simp add: 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
  2026
    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
  2027
  have eq2: "([(c',c)]\<bullet>M)[a\<turnstile>c>b] = ([(c',c)]\<bullet>(M[a\<turnstile>c>b]))" using fs2 a2 a1 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
  2028
    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
  2029
    apply(simp add: rename_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
  2030
    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
  2031
    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
  2032
  have "M{c:=(x).P}[a\<turnstile>c>b] = ([(c',c)]\<bullet>M){c':=(x).P}[a\<turnstile>c>b]" using eq by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(c',c)]\<bullet>M)[a\<turnstile>c>b]{c':=(x).P[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
  2034
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule substc_crename_comm)
9be4ab2acc13 split Class.thy into parts to 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(simp_all 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
  2038
    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
  2039
  also have "\<dots> = ([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P[a\<turnstile>c>b]}" using eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P}" using a2 by (simp add: 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
  2041
  also have "\<dots> = M[a\<turnstile>c>b]{c:=(x).P}" using eq' by simp
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  2044
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_nrename_comm':
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "x\<noteq>y" "x\<noteq>z" "y\<sharp>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
  2047
  shows "M{x:=<c>.P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{x:=<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
  2048
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
  2049
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
  2050
  assume a1: "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
  2051
  assume a1': "x\<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
  2052
  assume a2: "y\<sharp>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
  2053
  obtain x'::"name" where fs2: "x'\<sharp>(x,M,y,z)" by (rule exists_fresh(1), rule fin_supp, 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
  2054
  have eq: "M{x:=<c>.P} = ([(x',x)]\<bullet>M){x':=<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
  2055
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  2060
   have eq': "([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P} = M[y\<turnstile>n>z]{x:=<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
  2061
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule subst_rename[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
  2064
    apply(simp add: 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
  2065
    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
  2066
  have eq2: "([(x',x)]\<bullet>M)[y\<turnstile>n>z] = ([(x',x)]\<bullet>(M[y\<turnstile>n>z]))" using fs2 a2 a1 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
  2067
    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
  2068
    apply(simp add: rename_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
  2069
    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
  2070
    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
  2071
  have "M{x:=<c>.P}[y\<turnstile>n>z] = ([(x',x)]\<bullet>M){x':=<c>.P}[y\<turnstile>n>z]" using eq by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(x',x)]\<bullet>M)[y\<turnstile>n>z]{x':=<c>.P[y\<turnstile>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
  2073
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule substn_nrename_comm)
9be4ab2acc13 split Class.thy into parts to 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(simp_all 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
  2077
    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
  2078
  also have "\<dots> = ([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P[y\<turnstile>n>z]}" using eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P}" using a2 by (simp add: 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
  2080
  also have "\<dots> = M[y\<turnstile>n>z]{x:=<c>.P}" using eq' by simp
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  2083
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_nrename_comm':
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "x\<noteq>y" "y\<sharp>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
  2086
  shows "M{c:=(x).P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{c:=(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
  2087
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
  2088
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
  2089
  assume a1: "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
  2090
  assume a2: "y\<sharp>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
  2091
  obtain x'::"name" where fs2: "x'\<sharp>(x,P,y,z)" by (rule exists_fresh(1), rule fin_supp, 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
  2092
  have eq: "M{c:=(x).P} = M{c:=(x').([(x',x)]\<bullet>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
  2093
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  2098
   have eq': "M[y\<turnstile>n>z]{c:=(x).P} = M[y\<turnstile>n>z]{c:=(x').([(x',x)]\<bullet>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
  2099
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  2104
  have eq2: "([(x',x)]\<bullet>P)[y\<turnstile>n>z] = ([(x',x)]\<bullet>P)" using fs2 a2 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
  2105
    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
  2106
    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
  2107
    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
  2108
    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
  2109
  have "M{c:=(x).P}[y\<turnstile>n>z] = M{c:=(x').([(x',x)]\<bullet>P)}[y\<turnstile>n>z]" using eq by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = M[y\<turnstile>n>z]{c:=(x').(([(x',x)]\<bullet>P)[y\<turnstile>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
  2111
    using fs2
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule substc_nrename_comm)
9be4ab2acc13 split Class.thy into parts to 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(simp_all 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
  2115
    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
  2116
  also have "\<dots> = M[y\<turnstile>n>z]{c:=(x').(([(x',x)]\<bullet>P))}" using eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = M[y\<turnstile>n>z]{c:=(x).P}" using eq' by simp 
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  2120
9be4ab2acc13 split Class.thy into parts to 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
lemmas subst_comm = substn_crename_comm substc_crename_comm
9be4ab2acc13 split Class.thy into parts to 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
                    substn_nrename_comm substc_nrename_comm
9be4ab2acc13 split Class.thy into parts to 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
lemmas subst_comm' = substn_crename_comm' substc_crename_comm'
9be4ab2acc13 split Class.thy into parts to 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
                     substn_nrename_comm' substc_nrename_comm'
9be4ab2acc13 split Class.thy into parts to 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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  2126
text \<open>typing contexts\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
41798
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 39302
diff changeset
  2128
type_synonym ctxtn = "(name\<times>ty) list"
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 39302
diff changeset
  2129
type_synonym ctxtc = "(coname\<times>ty) 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
  2130
9be4ab2acc13 split Class.thy into parts to 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
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
  2132
  validc :: "ctxtc \<Rightarrow> bool"
9be4ab2acc13 split Class.thy into parts to 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
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
  2134
  vc1[intro]: "validc []"
9be4ab2acc13 split Class.thy into parts to 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
| vc2[intro]: "\<lbrakk>a\<sharp>\<Delta>; validc \<Delta>\<rbrakk> \<Longrightarrow> validc ((a,T)#\<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
  2136
9be4ab2acc13 split Class.thy into parts to 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
equivariance validc
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  2140
  validn :: "ctxtn \<Rightarrow> bool"
9be4ab2acc13 split Class.thy into parts to 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
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
  2142
  vn1[intro]: "validn []"
9be4ab2acc13 split Class.thy into parts to 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
| vn2[intro]: "\<lbrakk>x\<sharp>\<Gamma>; validn \<Gamma>\<rbrakk> \<Longrightarrow> validn ((x,T)#\<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
  2144
9be4ab2acc13 split Class.thy into parts to 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
equivariance validn
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  2148
  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
  2149
  and   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
  2150
  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
  2151
  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
  2152
  shows "a\<sharp>\<Gamma>" and "x\<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
  2153
proof -
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2154
  show "a\<sharp>\<Gamma>" by (induct \<Gamma>) (auto simp: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2155
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2156
  show "x\<sharp>\<Delta>" by (induct \<Delta>) (auto simp: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2158
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  2159
text \<open>cut-reductions\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
declare abs_perm[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
  2162
9be4ab2acc13 split Class.thy into parts to 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
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
  2164
  fin :: "trm \<Rightarrow> name \<Rightarrow> bool"
9be4ab2acc13 split Class.thy into parts to 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
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
  2166
  [intro]: "fin (Ax x a) 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
  2167
| [intro]: "x\<sharp>M \<Longrightarrow> fin (NotL <a>.M x) 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
  2168
| [intro]: "y\<sharp>[x].M \<Longrightarrow> fin (AndL1 (x).M y) 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
  2169
| [intro]: "y\<sharp>[x].M \<Longrightarrow> fin (AndL2 (x).M y) 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
  2170
| [intro]: "\<lbrakk>z\<sharp>[x].M;z\<sharp>[y].N\<rbrakk> \<Longrightarrow> fin (OrL (x).M (y).N z) 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
  2171
| [intro]: "\<lbrakk>y\<sharp>M;y\<sharp>[x].N\<rbrakk> \<Longrightarrow> fin (ImpL <a>.M (x).N y) 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
  2172
9be4ab2acc13 split Class.thy into parts to 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
equivariance fin
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_Ax_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
  2176
  assumes a: "fin (Ax x a) 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
  2177
  shows "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
  2178
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
  2179
apply(erule_tac fin.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2180
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2182
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_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
  2184
  assumes a: "fin (NotL <a>.M 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
  2185
  shows "x=y \<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
  2186
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
  2187
apply(erule_tac fin.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2188
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "y\<sharp>[aa].Ma")
9be4ab2acc13 split Class.thy into parts to 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(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
  2191
apply(simp_all 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
  2192
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
  2193
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_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
  2195
  assumes a: "fin (AndL1 (x).M 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
  2196
  shows "z=y \<and> z\<sharp>[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
  2197
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
  2198
apply(erule_tac fin.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2199
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2201
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_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
  2203
  assumes a: "fin (AndL2 (x).M 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
  2204
  shows "z=y \<and> z\<sharp>[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
  2205
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
  2206
apply(erule_tac fin.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2207
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2209
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_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
  2211
  assumes a: "fin (OrL (x).M (y).N u) 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
  2212
  shows "z=u \<and> z\<sharp>[x].M \<and> z\<sharp>[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
  2213
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
  2214
apply(erule_tac fin.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2215
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2217
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_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
  2219
  assumes a: "fin (ImpL <a>.M (x).N z) 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
  2220
  shows "z=y \<and> z\<sharp>M \<and> z\<sharp>[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
  2221
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
  2222
apply(erule_tac fin.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2223
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "y\<sharp>[aa].Ma")
9be4ab2acc13 split Class.thy into parts to 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(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
  2226
apply(simp_all add: abs_fresh)
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  2227
apply (metis abs_fresh(5))
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  2228
done
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  2229
 
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_rest_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
  2232
  shows "fin (Cut <a>.M (x).N) y \<Longrightarrow> 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
  2233
  and   "fin (NotR (x).M c) y \<Longrightarrow> 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
  2234
  and   "fin (AndR <a>.M <b>.N c) y \<Longrightarrow> 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
  2235
  and   "fin (OrR1 <a>.M b) y \<Longrightarrow> 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
  2236
  and   "fin (OrR2 <a>.M b) y \<Longrightarrow> 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
  2237
  and   "fin (ImpR (x).<a>.M b) y \<Longrightarrow> 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
  2238
by (erule fin.cases, 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
  2239
9be4ab2acc13 split Class.thy into parts to 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
lemmas fin_elims = fin_Ax_elim fin_NotL_elim fin_AndL1_elim fin_AndL2_elim fin_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
  2241
                   fin_ImpL_elim fin_rest_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
  2242
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_rename:
9be4ab2acc13 split Class.thy into parts to 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
  shows "fin M x \<Longrightarrow> fin ([(x',x)]\<bullet>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
  2245
by (induct rule: fin.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2246
   (auto simp: calc_atm simp add: fresh_left abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma not_fin_subst1:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "\<not>(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
  2250
  shows "\<not>(fin (M{c:=(y).P}) x)"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  2251
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(nominal_induct M avoiding: x c y P 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
  2253
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
  2254
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname:=(y).P},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
  2258
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
  2259
apply(simp add: 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
  2260
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
  2261
apply(simp add: fresh_fun_simp_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
  2262
apply(erule fin.cases, 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
  2263
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  2266
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
  2267
apply(drule freshn_after_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
  2268
apply(simp add: 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
  2269
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,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
  2270
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
  2271
apply(simp add: 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
  2272
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
  2273
apply(simp add: fresh_fun_simp_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
  2274
apply(erule fin.cases, 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
  2275
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,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
  2276
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
  2277
apply(simp add: 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
  2278
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
  2279
apply(simp add: fresh_fun_simp_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
  2280
apply(erule fin.cases, 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
  2281
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
apply(erule fin.cases, 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
  2283
apply(drule fin_AndL1_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2284
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(drule freshn_after_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
  2286
apply(subgoal_tac "name2\<sharp>[name1]. 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
  2287
apply(simp add: 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
  2288
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
  2289
apply(drule fin_AndL2_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2290
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(drule freshn_after_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
  2292
apply(subgoal_tac "name2\<sharp>[name1].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
  2293
apply(simp add: 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
  2294
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
  2295
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(y).P},coname1,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
  2296
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
  2297
apply(simp add: 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
  2298
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
  2299
apply(simp add: fresh_fun_simp_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
  2300
apply(erule fin.cases, 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
  2301
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(erule fin.cases, 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
  2303
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(y).P},coname1,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
  2304
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
  2305
apply(simp add: 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
  2306
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
  2307
apply(simp add: fresh_fun_simp_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
  2308
apply(erule fin.cases, 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
  2309
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
apply(erule fin.cases, 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
  2311
apply(drule fin_OrL_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2312
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(drule freshn_after_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
  2314
apply(subgoal_tac "name3\<sharp>[name1].trm1 \<and> name3\<sharp>[name2].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
  2315
apply(simp add: 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
  2316
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
  2317
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(y).P},coname1,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
  2318
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
  2319
apply(simp add: 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
  2320
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
  2321
apply(simp add: fresh_fun_simp_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
  2322
apply(erule fin.cases, 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
  2323
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
apply(erule fin.cases, 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
  2325
apply(drule fin_ImpL_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2326
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(drule freshn_after_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
  2328
apply(subgoal_tac "x\<sharp>[name1].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
  2329
apply(simp add: 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
  2330
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
  2331
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
  2332
9be4ab2acc13 split Class.thy into parts to 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
lemma not_fin_subst2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "\<not>(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
  2335
  shows "\<not>(fin (M{y:=<c>.P}) x)"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  2336
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(nominal_induct M avoiding: x c y P 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
  2338
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
  2339
apply(erule fin.cases, 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
  2340
apply(erule fin.cases, 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
  2341
apply(erule fin.cases, 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
  2342
apply(erule fin.cases, 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
  2343
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm{y:=<c>.P},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
  2344
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
  2345
apply(simp add: 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
  2346
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
  2347
apply(simp add: fresh_fun_simp_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
  2348
apply(erule fin.cases, 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
  2349
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(drule fin_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
  2351
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
  2352
apply(drule freshn_after_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
  2353
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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: 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
  2355
apply(erule fin.cases, 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
  2356
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm{y:=<c>.P},P,name1,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
  2357
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
  2358
apply(simp add: 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
  2359
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
  2360
apply(simp add: fresh_fun_simp_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
  2361
apply(erule fin.cases, 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
  2362
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_AndL1_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2364
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(drule freshn_after_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
  2366
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(subgoal_tac "name2\<sharp>[name1]. 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
  2368
apply(simp add: 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
  2369
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
  2370
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm{y:=<c>.P},P,name1,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
  2371
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
  2372
apply(simp add: 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
  2373
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
  2374
apply(simp add: fresh_fun_simp_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
  2375
apply(erule fin.cases, 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
  2376
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(drule fin_AndL2_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2378
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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 freshn_after_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
  2380
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "name2\<sharp>[name1].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
  2382
apply(simp add: 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
  2383
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
  2384
apply(erule fin.cases, 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
  2385
apply(erule fin.cases, 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
  2386
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm1{y:=<c>.P},trm2{y:=<c>.P},name1,name2,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
  2387
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
  2388
apply(simp add: 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
  2389
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
  2390
apply(simp add: fresh_fun_simp_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
  2391
apply(erule fin.cases, 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
  2392
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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 fin_OrL_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2394
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(drule freshn_after_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
  2396
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule freshn_after_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
  2398
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "name3\<sharp>[name1].trm1 \<and> name3\<sharp>[name2].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
  2400
apply(simp add: 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
  2401
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
  2402
apply(erule fin.cases, 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
  2403
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},name1,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
  2404
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
  2405
apply(simp add: 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
  2406
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
  2407
apply(simp add: fresh_fun_simp_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
  2408
apply(erule fin.cases, 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
  2409
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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 fin_ImpL_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2411
apply(auto simp: abs_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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 freshn_after_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
  2413
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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 freshn_after_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
  2415
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "x\<sharp>[name1].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
  2417
apply(simp add: 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
  2418
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
  2419
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
  2420
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_subst1:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "fin M x" "x\<noteq>y" "x\<sharp>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
  2423
  shows "fin (M{y:=<c>.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
  2424
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
  2425
apply(nominal_induct M avoiding: x y c P 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
  2426
apply(auto dest!: fin_elims simp add: subst_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
  2427
apply(rule fin.intros, simp add: subst_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
  2428
apply(rule fin.intros, simp add: subst_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
  2429
apply(rule fin.intros, simp add: subst_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
  2430
apply(rule fin.intros, simp add: subst_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
  2431
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_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
  2432
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_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
  2433
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_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
  2434
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_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
  2435
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_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
  2436
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
  2437
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_subst2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "fin M y" "x\<noteq>y" "y\<sharp>P" "M\<noteq>Ax 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
  2440
  shows "fin (M{c:=(x).P}) 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
  2441
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
  2442
apply(nominal_induct M avoiding: x y c P 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
  2443
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
  2444
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
  2445
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
  2446
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  2450
apply(rule 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
  2451
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2455
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
  2456
apply(rule 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
  2457
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
  2458
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2460
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
  2461
apply(rule 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
  2462
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
  2463
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2467
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
  2468
apply(rule 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
  2469
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
  2470
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
  2471
apply(rule 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
  2472
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
  2473
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2476
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
  2477
apply(rule 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
  2478
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
  2479
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
  2480
apply(rule 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
  2481
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
  2482
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
  2483
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_substn_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
  2485
  assumes a: "fin M x" "x\<noteq>y" "x\<sharp>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
  2486
  shows "M[x\<turnstile>n>y]{y:=<c>.P} = Cut <c>.P (x).(M{y:=<c>.P})"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  2487
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(nominal_induct M avoiding: x y c P 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
  2489
apply(drule fin_Ax_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
  2490
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2492
apply(simp add: alpha 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
  2493
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_rest_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
  2495
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_rest_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
  2497
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_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
  2499
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(subgoal_tac "\<exists>z::name. z\<sharp>(trm,y,x,P,trm[x\<turnstile>n>y]{y:=<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
  2501
apply(erule exE, simp add: 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
  2502
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
  2503
apply(simp add: fresh_fun_simp_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
  2504
apply(simp add: trm.inject alpha fresh_atm calc_atm 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
  2505
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
  2506
apply(simp add: nsubst_eqvt 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
  2507
apply(simp add: perm_fresh_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
  2508
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
  2509
apply(rule 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
  2510
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_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
  2512
apply(rule 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
  2513
apply(drule fin_rest_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
  2514
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_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
  2516
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>z::name. z\<sharp>(name2,name1,P,trm[name2\<turnstile>n>y]{y:=<c>.P},y,P,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
  2518
apply(erule exE, simp add: 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
  2519
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
  2520
apply(simp add: fresh_fun_simp_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
  2521
apply(simp add: trm.inject alpha fresh_atm calc_atm 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
  2522
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
  2523
apply(simp add: nsubst_eqvt 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
  2524
apply(simp add: perm_fresh_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
  2525
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
  2526
apply(rule 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
  2527
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_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
  2529
apply(rule 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
  2530
apply(drule fin_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
  2531
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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 "\<exists>z::name. z\<sharp>(name2,name1,P,trm[name2\<turnstile>n>y]{y:=<c>.P},y,P,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
  2533
apply(erule exE, simp add: 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
  2534
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
  2535
apply(simp add: fresh_fun_simp_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
  2536
apply(simp add: trm.inject alpha fresh_atm calc_atm 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
  2537
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
  2538
apply(simp add: nsubst_eqvt 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
  2539
apply(simp add: perm_fresh_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
  2540
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
  2541
apply(rule 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
  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
apply(rule exists_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
  2544
apply(rule 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
  2545
apply(drule fin_rest_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
  2546
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_rest_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
  2548
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_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
  2550
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
  2551
apply(simp add: subst_fresh 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
  2552
apply(subgoal_tac "\<exists>z::name. z\<sharp>(name3,name2,name1,P,trm1[name3\<turnstile>n>y]{y:=<c>.P},trm2[name3\<turnstile>n>y]{y:=<c>.P},y,P,trm1,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
  2553
apply(erule exE, simp add: 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
  2554
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
  2555
apply(simp add: fresh_fun_simp_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
  2556
apply(simp add: trm.inject alpha fresh_atm calc_atm 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
  2557
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
  2558
apply(simp add: nsubst_eqvt 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
  2559
apply(simp add: perm_fresh_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
  2560
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
  2561
apply(simp add: nsubst_eqvt 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
  2562
apply(simp add: perm_fresh_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
  2563
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
  2564
apply(rule exists_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
  2565
apply(rule 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
  2566
apply(drule fin_rest_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
  2567
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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 fin_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
  2569
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
  2570
apply(simp add: subst_fresh 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
  2571
apply(subgoal_tac "\<exists>z::name. z\<sharp>(name1,x,P,trm1[x\<turnstile>n>y]{y:=<c>.P},trm2[x\<turnstile>n>y]{y:=<c>.P},y,P,trm1,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
  2572
apply(erule exE, simp add: 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
  2573
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
  2574
apply(simp add: fresh_fun_simp_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
  2575
apply(simp add: trm.inject alpha fresh_atm calc_atm 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
  2576
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
  2577
apply(simp add: nsubst_eqvt 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
  2578
apply(simp add: perm_fresh_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
  2579
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
  2580
apply(simp add: nsubst_eqvt 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
  2581
apply(simp add: perm_fresh_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
  2582
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
  2583
apply(rule exists_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
  2584
apply(rule 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
  2585
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
  2586
9be4ab2acc13 split Class.thy into parts to 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
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
  2588
  fic :: "trm \<Rightarrow> coname \<Rightarrow> bool"
9be4ab2acc13 split Class.thy into parts to 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
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
  2590
  [intro]: "fic (Ax x a) 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
  2591
| [intro]: "a\<sharp>M \<Longrightarrow> fic (NotR (x).M a) 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
  2592
| [intro]: "\<lbrakk>c\<sharp>[a].M;c\<sharp>[b].N\<rbrakk> \<Longrightarrow> fic (AndR <a>.M <b>.N 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
  2593
| [intro]: "b\<sharp>[a].M \<Longrightarrow> fic (OrR1 <a>.M 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
  2594
| [intro]: "b\<sharp>[a].M \<Longrightarrow> fic (OrR2 <a>.M 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
  2595
| [intro]: "\<lbrakk>b\<sharp>[a].M\<rbrakk> \<Longrightarrow> fic (ImpR (x).<a>.M 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
  2596
9be4ab2acc13 split Class.thy into parts to 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
equivariance fic
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_Ax_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
  2600
  assumes a: "fic (Ax x 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
  2601
  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
  2602
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
  2603
apply(erule_tac fic.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2604
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2606
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_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
  2608
  assumes a: "fic (NotR (x).M 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
  2609
  shows "a=b \<and> b\<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
  2610
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
  2611
apply(erule_tac fic.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2612
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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 "b\<sharp>[xa].Ma")
9be4ab2acc13 split Class.thy into parts to 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
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
  2615
apply(simp_all 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
  2616
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
  2617
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_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
  2619
  assumes a: "fic (OrR1 <a>.M 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
  2620
  shows "b=c \<and> c\<sharp>[a].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
  2621
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
  2622
apply(erule_tac fic.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2623
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2625
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_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
  2627
  assumes a: "fic (OrR2 <a>.M 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
  2628
  shows "b=c \<and> c\<sharp>[a].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
  2629
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
  2630
apply(erule_tac fic.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2631
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2633
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_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
  2635
  assumes a: "fic (AndR <a>.M <b>.N c) 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
  2636
  shows "c=d \<and> d\<sharp>[a].M \<and> d\<sharp>[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
  2637
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
  2638
apply(erule_tac fic.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2639
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  2641
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_ImpR_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
  2643
  assumes a: "fic (ImpR (x).<a>.M 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
  2644
  shows "b=c \<and> b\<sharp>[a].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
  2645
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
  2646
apply(erule_tac fic.cases)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2647
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "c\<sharp>[xa].[aa].Ma")
9be4ab2acc13 split Class.thy into parts to 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 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
  2650
apply(simp_all 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
  2651
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
  2652
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_rest_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
  2654
  shows "fic (Cut <a>.M (x).N) d \<Longrightarrow> 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
  2655
  and   "fic (NotL <a>.M x) d \<Longrightarrow> 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
  2656
  and   "fic (OrL (x).M (y).N z) d \<Longrightarrow> 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
  2657
  and   "fic (AndL1 (x).M y) d \<Longrightarrow> 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
  2658
  and   "fic (AndL2 (x).M y) d \<Longrightarrow> 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
  2659
  and   "fic (ImpL <a>.M (x).N y) d \<Longrightarrow> 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
  2660
by (erule fic.cases, 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
  2661
9be4ab2acc13 split Class.thy into parts to 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
lemmas fic_elims = fic_Ax_elim fic_NotR_elim fic_OrR1_elim fic_OrR2_elim fic_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
  2663
                   fic_ImpR_elim fic_rest_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
  2664
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_rename:
9be4ab2acc13 split Class.thy into parts to 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
  shows "fic M a \<Longrightarrow> fic ([(a',a)]\<bullet>M) 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
  2667
by (induct rule: fic.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2668
   (auto simp: calc_atm simp add: fresh_left abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma not_fic_subst1:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "\<not>(fic M 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
  2672
  shows "\<not>(fic (M{y:=<c>.P}) 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
  2673
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
  2674
apply(nominal_induct M avoiding: a c y P 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
  2675
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
  2676
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2680
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
  2681
apply(drule freshc_after_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
  2682
apply(simp add: 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
  2683
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,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
  2684
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
  2685
apply(simp add: 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
  2686
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
  2687
apply(simp add: fresh_fun_simp_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
  2688
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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 fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2692
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
  2693
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
  2694
apply(drule freshc_after_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
  2695
apply(drule freshc_after_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
  2696
apply(simp add: fic.intros 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
  2697
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1,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
  2698
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
  2699
apply(simp add: 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
  2700
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
  2701
apply(simp add: fresh_fun_simp_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
  2702
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1,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
  2706
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
  2707
apply(simp add: 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
  2708
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
  2709
apply(simp add: fresh_fun_simp_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
  2710
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2714
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
  2715
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
  2716
apply(drule freshc_after_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
  2717
apply(simp add: fic.intros 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
  2718
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
  2719
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
  2720
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
  2721
apply(drule freshc_after_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
  2722
apply(simp add: fic.intros 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
  2723
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},trm2{y:=<c>.P},P,name1,name2,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
  2724
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
  2725
apply(simp add: 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
  2726
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
  2727
apply(simp add: fresh_fun_simp_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
  2728
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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 fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2733
apply(drule freshc_after_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
  2734
apply(simp add: fic.intros 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
  2735
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},P,name1,name2,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
  2736
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
  2737
apply(simp add: 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
  2738
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
  2739
apply(simp add: fresh_fun_simp_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
  2740
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  2744
9be4ab2acc13 split Class.thy into parts to 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
lemma not_fic_subst2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "\<not>(fic M 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
  2747
  shows "\<not>(fic (M{c:=(y).P}) 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
  2748
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
  2749
apply(nominal_induct M avoiding: a c y P 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
  2750
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
  2751
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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 "\<exists>c'::coname. c'\<sharp>(trm{coname:=(y).P},P,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
  2755
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
  2756
apply(simp add: 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
  2757
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
  2758
apply(simp add: fresh_fun_simp_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
  2759
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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 exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule freshc_after_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
  2763
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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 add: fic.intros 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
  2765
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,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
  2767
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
  2768
apply(simp add: 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
  2769
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
  2770
apply(simp add: fresh_fun_simp_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
  2771
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2775
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
  2776
apply(drule freshc_after_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
  2777
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2778
apply(drule freshc_after_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
  2779
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: fic.intros 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
  2781
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(y).P},P,coname1,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
  2784
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
  2785
apply(simp add: 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
  2786
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
  2787
apply(simp add: fresh_fun_simp_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
  2788
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2792
apply(drule freshc_after_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
  2793
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: fic.intros 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
  2795
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(y).P},P,coname1,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
  2796
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
  2797
apply(simp add: 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
  2798
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
  2799
apply(simp add: fresh_fun_simp_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
  2800
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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 fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2804
apply(drule freshc_after_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
  2805
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: fic.intros 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
  2807
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(y).P},P,coname1,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
  2809
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
  2810
apply(simp add: 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
  2811
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
  2812
apply(simp add: fresh_fun_simp_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
  2813
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2817
apply(drule freshc_after_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
  2818
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: fic.intros 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
  2820
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  2822
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_subst1:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "fic M a" "a\<noteq>b" "a\<sharp>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
  2825
  shows "fic (M{b:=(x).P}) 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
  2826
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
  2827
apply(nominal_induct M avoiding: x b a P 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
  2828
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
  2829
apply(simp add: 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
  2830
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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 fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2833
apply(rule 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
  2834
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2838
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
  2839
apply(rule 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
  2840
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
  2841
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
  2842
apply(rule 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
  2843
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
  2844
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2848
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
  2849
apply(rule 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
  2850
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
  2851
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2853
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
  2854
apply(rule 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
  2855
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
  2856
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2859
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
  2860
apply(rule 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
  2861
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
  2862
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  2864
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_subst2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "fic M a" "c\<noteq>a" "a\<sharp>P" "M\<noteq>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
  2867
  shows "fic (M{x:=<c>.P}) 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
  2868
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
  2869
apply(nominal_induct M avoiding: x a c P 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
  2870
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
  2871
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
  2872
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
  2873
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2876
apply(rule 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
  2877
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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 fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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 fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2881
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
  2882
apply(rule 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
  2883
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
  2884
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
  2885
apply(rule 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
  2886
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
  2887
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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 fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2891
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
  2892
apply(rule 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
  2893
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
  2894
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2896
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
  2897
apply(rule 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
  2898
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
  2899
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2902
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
  2903
apply(rule 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
  2904
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
  2905
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  2907
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_substc_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
  2909
  assumes a: "fic M a" "a\<noteq>b" "a\<sharp>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
  2910
  shows "M[a\<turnstile>c>b]{b:=(y).P} = Cut <a>.(M{b:=(y).P}) (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
  2911
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
  2912
apply(nominal_induct M avoiding: a b  y P 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
  2913
apply(drule fic_Ax_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
  2914
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2916
apply(simp add: alpha calc_atm fresh_atm 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
  2917
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2918
apply(drule fic_rest_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
  2919
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_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
  2921
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2923
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
  2924
apply(simp add: trm.inject alpha fresh_atm fresh_prod fresh_atm calc_atm 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(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
  2926
apply(simp add: csubst_eqvt 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
  2927
apply(simp add: perm_fresh_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
  2928
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
  2929
apply(rule 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
  2930
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_rest_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
  2932
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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 fic_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
  2934
apply(simp add: abs_fresh fresh_atm subst_fresh 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
  2935
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
  2936
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
  2937
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh 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
  2938
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
  2939
apply(simp add: csubst_eqvt 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
  2940
apply(simp add: perm_fresh_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
  2941
apply(simp add: csubst_eqvt 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
  2942
apply(simp add: perm_fresh_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
  2943
apply(simp add: 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
  2944
apply(drule fic_rest_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
  2945
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_rest_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
  2947
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_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
  2949
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  2951
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
  2952
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh 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
  2953
apply(simp add: csubst_eqvt 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
  2954
apply(simp add: perm_fresh_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
  2955
apply(simp add: subst_fresh 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
  2956
apply(drule fic_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
  2957
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
  2958
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
  2959
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
  2960
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh 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
  2961
apply(simp add: csubst_eqvt 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
  2962
apply(simp add: perm_fresh_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
  2963
apply(simp add: subst_fresh 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
  2964
apply(drule fic_rest_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
  2965
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_ImpR_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
  2967
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
  2968
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
  2969
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
  2970
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh 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
  2971
apply(simp add: csubst_eqvt 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
  2972
apply(simp add: perm_fresh_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
  2973
apply(simp add: subst_fresh 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
  2974
apply(drule fic_rest_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
  2975
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  2977
9be4ab2acc13 split Class.thy into parts to 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
inductive
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2979
  l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>l _" [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to 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
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2981
  LAxR:  "\<lbrakk>x\<sharp>M; a\<sharp>b; fic M a\<rbrakk> \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^sub>l M[a\<turnstile>c>b]"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2982
| LAxL:  "\<lbrakk>a\<sharp>M; x\<sharp>y; fin M x\<rbrakk> \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^sub>l M[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
  2983
| LNot:  "\<lbrakk>y\<sharp>(M,N); x\<sharp>(N,y); a\<sharp>(M,N,b); b\<sharp>M; y\<noteq>x; b\<noteq>a\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2984
          Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^sub>l Cut <b>.N (x).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
  2985
| LAnd1: "\<lbrakk>b\<sharp>([a1].M1,[a2].M2,N,a1,a2); y\<sharp>([x].N,M1,M2,x); x\<sharp>(M1,M2); a1\<sharp>(M2,N); a2\<sharp>(M1,N); a1\<noteq>a2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2986
          Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^sub>l Cut <a1>.M1 (x).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
  2987
| LAnd2: "\<lbrakk>b\<sharp>([a1].M1,[a2].M2,N,a1,a2); y\<sharp>([x].N,M1,M2,x); x\<sharp>(M1,M2); a1\<sharp>(M2,N); a2\<sharp>(M1,N); a1\<noteq>a2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2988
          Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL2 (x).N y) \<longrightarrow>\<^sub>l Cut <a2>.M2 (x).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
  2989
| LOr1:  "\<lbrakk>b\<sharp>([a].M,N1,N2,a); y\<sharp>([x1].N1,[x2].N2,M,x1,x2); x1\<sharp>(M,N2); x2\<sharp>(M,N1); a\<sharp>(N1,N2); x1\<noteq>x2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2990
          Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x1).N1"
36277
9be4ab2acc13 split Class.thy into parts to 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
| LOr2:  "\<lbrakk>b\<sharp>([a].M,N1,N2,a); y\<sharp>([x1].N1,[x2].N2,M,x1,x2); x1\<sharp>(M,N2); x2\<sharp>(M,N1); a\<sharp>(N1,N2); x1\<noteq>x2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2992
          Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x2).N2"
36277
9be4ab2acc13 split Class.thy into parts to 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
| LImp:  "\<lbrakk>z\<sharp>(N,[y].P,[x].M,y,x); b\<sharp>([a].M,[c].N,P,c,a); x\<sharp>(N,[y].P,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
  2994
          c\<sharp>(P,[a].M,b,a); a\<sharp>([c].N,P); y\<sharp>(N,[x].M)\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2995
          Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^sub>l Cut <a>.(Cut <c>.N (x).M) (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
  2996
9be4ab2acc13 split Class.thy into parts to 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
equivariance l_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
  2998
9be4ab2acc13 split Class.thy into parts to 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
lemma l_redu_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
  3000
  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
  3001
  and   pi2::"coname prm"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3002
  shows "(pi1\<bullet>M) \<longrightarrow>\<^sub>l (pi1\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^sub>l M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3003
  and   "(pi2\<bullet>M) \<longrightarrow>\<^sub>l (pi2\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^sub>l 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
  3004
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
  3005
apply(drule_tac pi="rev pi1" in l_redu.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
  3006
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
  3007
apply(drule_tac pi="rev pi2" in l_redu.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
  3008
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
  3009
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
  3010
9be4ab2acc13 split Class.thy into parts to 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
nominal_inductive l_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
  3012
  apply(simp_all add: abs_fresh fresh_atm rename_fresh 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
  3013
  apply(force)+
9be4ab2acc13 split Class.thy into parts to 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
  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
  3015
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_l_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
  3017
  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
  3018
  and   a::"coname"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3019
  shows "M \<longrightarrow>\<^sub>l M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3020
  and   "M \<longrightarrow>\<^sub>l M' \<Longrightarrow> a\<sharp>M \<Longrightarrow> a\<sharp>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
  3021
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
  3022
apply(induct rule: l_redu.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3023
apply(auto simp: abs_fresh rename_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(case_tac "xa=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
  3025
apply(simp add: 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
  3026
apply(simp add: rename_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
  3027
apply(simp add: fresh_prod 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
  3028
apply(induct rule: l_redu.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3029
apply(auto simp: abs_fresh rename_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(case_tac "aa=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
  3031
apply(simp add: 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
  3032
apply(simp add: rename_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
  3033
apply(simp add: fresh_prod 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
  3034
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
  3035
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LAxR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3037
  shows "fic M a \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^sub>l M[a\<turnstile>c>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
  3038
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
  3039
  assume fin: "fic M 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
  3040
  obtain x'::"name" where fs1: "x'\<sharp>(M,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3041
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M,b)" by (rule exists_fresh(2), rule fin_supp, 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
  3042
  have "Cut <a>.M (x).(Ax x b) =  Cut <a'>.([(a',a)]\<bullet>M) (x').(Ax x' b)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3043
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3044
  also have "\<dots> \<longrightarrow>\<^sub>l ([(a',a)]\<bullet>M)[a'\<turnstile>c>b]" using fs1 fs2 fin
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: l_redu.intros simp add: fresh_left calc_atm fic_rename)
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = M[a\<turnstile>c>b]" using fs1 fs2 by (simp add: crename_rename)
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3049
    
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LAxL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3051
  shows "fin M x \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^sub>l M[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
  3052
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
  3053
  assume fin: "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
  3054
  obtain x'::"name" where fs1: "x'\<sharp>(y,M,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3055
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M)" by (rule exists_fresh(2), rule fin_supp, 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
  3056
  have "Cut <a>.(Ax y a) (x).M = Cut <a'>.(Ax y a') (x').([(x',x)]\<bullet>M)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3057
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3058
  also have "\<dots> \<longrightarrow>\<^sub>l ([(x',x)]\<bullet>M)[x'\<turnstile>n>y]" using fs1 fs2 fin
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: l_redu.intros simp add: fresh_left calc_atm fin_rename)
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = M[x\<turnstile>n>y]" using fs1 fs2 by (simp add: nrename_rename)
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3063
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LNot_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3065
  shows "\<lbrakk>y\<sharp>N; a\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^sub>l Cut <b>.N (x).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
  3066
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
  3067
  assume fs: "y\<sharp>N" "a\<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
  3068
  obtain x'::"name" where f1: "x'\<sharp>(y,N,M,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3069
  obtain y'::"name" where f2: "y'\<sharp>(y,N,M,x,x')" by (rule exists_fresh(1), rule fin_supp, 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
  3070
  obtain a'::"coname" where f3: "a'\<sharp>(a,M,N,b)" by (rule exists_fresh(2), rule fin_supp, 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
  3071
  obtain b'::"coname" where f4: "b'\<sharp>(a,M,N,b,a')" by (rule exists_fresh(2), rule fin_supp, 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
  3072
  have "Cut <a>.(NotR (x).M a) (y).(NotL <b>.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
  3073
                      = Cut <a'>.(NotR (x).([(a',a)]\<bullet>M) a') (y').(NotL <b>.([(y',y)]\<bullet>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
  3074
    using f1 f2 f3 f4 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3075
    by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = Cut <a'>.(NotR (x).M a') (y').(NotL <b>.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
  3077
    using f1 f2 f3 f4 fs by (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
  3078
  also have "\<dots> = Cut <a'>.(NotR (x').([(x',x)]\<bullet>M) a') (y').(NotL <b'>.([(b',b)]\<bullet>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
  3079
    using f1 f2 f3 f4 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3080
    by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3081
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <b'>.([(b',b)]\<bullet>N) (x').([(x',x)]\<bullet>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
  3082
    using f1 f2 f3 f4 fs
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro:  l_redu.intros simp add: fresh_prod fresh_left 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
  3084
  also have "\<dots> = Cut <b>.N (x).M"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3085
    using f1 f2 f3 f4 by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3088
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LAnd1_intro[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
  3090
  shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3091
         \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).N y) \<longrightarrow>\<^sub>l Cut <b1>.M1 (x).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
  3092
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
  3093
  assume fs: "a\<sharp>([b1].M1,[b2].M2)" "y\<sharp>[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
  3094
  obtain x'::"name" where f1: "x'\<sharp>(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3095
  obtain y'::"name" where f2: "y'\<sharp>(y,N,M1,M2,x,x')" by (rule exists_fresh(1), rule fin_supp, 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
  3096
  obtain a'::"coname" where f3: "a'\<sharp>(a,M1,M2,N,b1,b2)" by (rule exists_fresh(2), rule fin_supp, 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
  3097
  obtain b1'::"coname" where f4:"b1'\<sharp>(a,M1,M2,N,b1,b2,a')" by (rule exists_fresh(2), rule fin_supp, 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
  3098
  obtain b2'::"coname" where f5:"b2'\<sharp>(a,M1,M2,N,b1,b2,a',b1')" by (rule exists_fresh(2),rule fin_supp, 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
  3099
  have "Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).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
  3100
                      = Cut <a'>.(AndR <b1>.M1 <b2>.M2 a') (y').(AndL1 (x).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
  3101
    using f1 f2 f3 f4 fs
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3102
    apply(rule_tac 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
  3103
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3104
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3106
  also have "\<dots> = Cut <a'>.(AndR <b1'>.([(b1',b1)]\<bullet>M1) <b2'>.([(b2',b2)]\<bullet>M2) 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
  3107
                                                               (y').(AndL1 (x').([(x',x)]\<bullet>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
  3108
    using f1 f2 f3 f4 f5 fs 
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3110
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm 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
  3111
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3112
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <b1'>.([(b1',b1)]\<bullet>M1) (x').([(x',x)]\<bullet>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
  3113
    using f1 f2 f3 f4 f5 fs
9be4ab2acc13 split Class.thy into parts to 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
    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
  3115
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3116
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3118
  also have "\<dots> = Cut <b1>.M1 (x).N"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3119
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3122
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3123
lemma better_LAnd2_intro[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
  3124
  shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3125
         \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).N y) \<longrightarrow>\<^sub>l Cut <b2>.M2 (x).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
  3126
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
  3127
  assume fs: "a\<sharp>([b1].M1,[b2].M2)" "y\<sharp>[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
  3128
  obtain x'::"name" where f1: "x'\<sharp>(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3129
  obtain y'::"name" where f2: "y'\<sharp>(y,N,M1,M2,x,x')" by (rule exists_fresh(1), rule fin_supp, 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
  3130
  obtain a'::"coname" where f3: "a'\<sharp>(a,M1,M2,N,b1,b2)" by (rule exists_fresh(2), rule fin_supp, 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
  3131
  obtain b1'::"coname" where f4:"b1'\<sharp>(a,M1,M2,N,b1,b2,a')" by (rule exists_fresh(2), rule fin_supp, 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
  3132
  obtain b2'::"coname" where f5:"b2'\<sharp>(a,M1,M2,N,b1,b2,a',b1')" by (rule exists_fresh(2),rule fin_supp, 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
  3133
  have "Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).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
  3134
                      = Cut <a'>.(AndR <b1>.M1 <b2>.M2 a') (y').(AndL2 (x).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
  3135
    using f1 f2 f3 f4 fs
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3137
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3138
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3140
  also have "\<dots> = Cut <a'>.(AndR <b1'>.([(b1',b1)]\<bullet>M1) <b2'>.([(b2',b2)]\<bullet>M2) 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
  3141
                                                               (y').(AndL2 (x').([(x',x)]\<bullet>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
  3142
    using f1 f2 f3 f4 f5 fs 
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3144
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm 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
  3145
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3146
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <b2'>.([(b2',b2)]\<bullet>M2) (x').([(x',x)]\<bullet>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
  3147
    using f1 f2 f3 f4 f5 fs
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3148
    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
  3149
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3150
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3152
  also have "\<dots> = Cut <b2>.M2 (x).N"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3153
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  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 better_LOr1_intro[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
  3158
  shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3159
         \<Longrightarrow> Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x1).N1"
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3161
  assume fs: "y\<sharp>([x1].N1,[x2].N2)" "b\<sharp>[a].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
  3162
  obtain y'::"name" where f1: "y'\<sharp>(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, 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
  3163
  obtain x1'::"name" where f2: "x1'\<sharp>(y,M,N1,N2,x1,x2,y')" by (rule exists_fresh(1), rule fin_supp, 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
  3164
  obtain x2'::"name" where f3: "x2'\<sharp>(y,M,N1,N2,x1,x2,y',x1')" by (rule exists_fresh(1), rule fin_supp, 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
  3165
  obtain a'::"coname" where f4: "a'\<sharp>(a,N1,N2,M,b)" by (rule exists_fresh(2), rule fin_supp, 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
  3166
  obtain b'::"coname" where f5: "b'\<sharp>(a,N1,N2,M,b,a')" by (rule exists_fresh(2),rule fin_supp, 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
  3167
  have "Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 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
  3168
                      = Cut <b'>.(OrR1 <a>.M b') (y').(OrL (x1).N1 (x2).N2 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
  3169
    using f1 f2 f3 f4 f5 fs
9be4ab2acc13 split Class.thy into parts to 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(rule_tac 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
  3171
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3172
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3174
  also have "\<dots> = Cut <b'>.(OrR1 <a'>.([(a',a)]\<bullet>M) 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
  3175
              (y').(OrL (x1').([(x1',x1)]\<bullet>N1) (x2').([(x2',x2)]\<bullet>N2) 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
  3176
    using f1 f2 f3 f4 f5 fs 
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3178
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm 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
  3179
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3180
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.([(a',a)]\<bullet>M) (x1').([(x1',x1)]\<bullet>N1)"
36277
9be4ab2acc13 split Class.thy into parts to 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
    using f1 f2 f3 f4 f5 fs
9be4ab2acc13 split Class.thy into parts to 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
    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
  3183
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3184
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3186
  also have "\<dots> = Cut <a>.M (x1).N1"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3187
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3188
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3189
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
  3190
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3191
lemma better_LOr2_intro[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
  3192
  shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3193
         \<Longrightarrow> Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x2).N2"
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3195
  assume fs: "y\<sharp>([x1].N1,[x2].N2)" "b\<sharp>[a].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
  3196
  obtain y'::"name" where f1: "y'\<sharp>(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, 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
  3197
  obtain x1'::"name" where f2: "x1'\<sharp>(y,M,N1,N2,x1,x2,y')" by (rule exists_fresh(1), rule fin_supp, 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
  3198
  obtain x2'::"name" where f3: "x2'\<sharp>(y,M,N1,N2,x1,x2,y',x1')" by (rule exists_fresh(1), rule fin_supp, 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
  3199
  obtain a'::"coname" where f4: "a'\<sharp>(a,N1,N2,M,b)" by (rule exists_fresh(2), rule fin_supp, 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
  3200
  obtain b'::"coname" where f5: "b'\<sharp>(a,N1,N2,M,b,a')" by (rule exists_fresh(2),rule fin_supp, 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
  3201
  have "Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 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
  3202
                      = Cut <b'>.(OrR2 <a>.M b') (y').(OrL (x1).N1 (x2).N2 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
  3203
    using f1 f2 f3 f4 f5 fs
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3205
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3206
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3208
  also have "\<dots> = Cut <b'>.(OrR2 <a'>.([(a',a)]\<bullet>M) 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
  3209
              (y').(OrL (x1').([(x1',x1)]\<bullet>N1) (x2').([(x2',x2)]\<bullet>N2) 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
  3210
    using f1 f2 f3 f4 f5 fs 
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3212
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm 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
  3213
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3214
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.([(a',a)]\<bullet>M) (x2').([(x2',x2)]\<bullet>N2)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3215
    using f1 f2 f3 f4 f5 fs
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3217
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3218
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3219
    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
  3220
  also have "\<dots> = Cut <a>.M (x2).N2"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3221
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3224
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LImp_intro[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
  3226
  shows "\<lbrakk>z\<sharp>(N,[y].P); b\<sharp>[a].M; a\<sharp>N\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3227
         \<Longrightarrow> Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^sub>l Cut <a>.(Cut <c>.N (x).M) (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
  3228
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
  3229
  assume fs: "z\<sharp>(N,[y].P)" "b\<sharp>[a].M" "a\<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
  3230
  obtain y'::"name" where f1: "y'\<sharp>(y,M,N,P,z,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3231
  obtain x'::"name" where f2: "x'\<sharp>(y,M,N,P,z,x,y')" by (rule exists_fresh(1), rule fin_supp, 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
  3232
  obtain z'::"name" where f3: "z'\<sharp>(y,M,N,P,z,x,y',x')" by (rule exists_fresh(1), rule fin_supp, 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
  3233
  obtain a'::"coname" where f4: "a'\<sharp>(a,N,P,M,b)" by (rule exists_fresh(2), rule fin_supp, 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
  3234
  obtain b'::"coname" where f5: "b'\<sharp>(a,N,P,M,b,c,a')" by (rule exists_fresh(2),rule fin_supp, 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
  3235
  obtain c'::"coname" where f6: "c'\<sharp>(a,N,P,M,b,c,a',b')" by (rule exists_fresh(2),rule fin_supp, 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
  3236
  have " Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P 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
  3237
                      =  Cut <b'>.(ImpR (x).<a>.M b') (z').(ImpL <c>.N (y).P 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
  3238
    using f1 f2 f3 f4 f5 fs
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3240
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3241
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    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
  3243
  also have "\<dots> = Cut <b'>.(ImpR (x').<a'>.([(a',a)]\<bullet>([(x',x)]\<bullet>M)) 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
  3244
                           (z').(ImpL <c'>.([(c',c)]\<bullet>N) (y').([(y',y)]\<bullet>P) 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
  3245
    using f1 f2 f3 f4 f5 f6 fs 
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule_tac 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
  3247
    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
  3248
    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
  3249
    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
  3250
    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
  3251
    apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left 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
  3252
    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
  3253
    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
  3254
    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
  3255
    apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left 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
  3256
    apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left 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
  3257
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3258
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.(Cut <c'>.([(c',c)]\<bullet>N) (x').([(a',a)]\<bullet>[(x',x)]\<bullet>M)) (y').([(y',y)]\<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
  3259
    using f1 f2 f3 f4 f5 f6 fs
9be4ab2acc13 split Class.thy into parts to 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
    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
  3261
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3262
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3263
    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
  3264
  also have "\<dots> = Cut <a>.(Cut <c>.N (x).M) (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
  3265
    using f1 f2 f3 f4 f5 f6 fs 
9be4ab2acc13 split Class.thy into parts to 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
    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
  3267
    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
  3268
    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
  3269
    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
  3270
    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
  3271
    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
  3272
    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
  3273
    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
  3274
    apply(perm_simp add: calc_atm)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3275
    apply(auto simp: fresh_prod fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(perm_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
  3277
    apply(perm_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
  3278
    apply(perm_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
  3279
    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
  3280
    apply(perm_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
  3281
    apply(rule_tac pi="[(a',a)]" in pt_bij4[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
  3282
    apply(perm_simp add: 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
  3283
    apply(perm_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
  3284
    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
  3285
    apply(perm_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
  3286
    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
  3287
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  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 alpha_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
  3291
  fixes M::"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
  3292
  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
  3293
  assumes a: "[a].M = [b].N" "c\<sharp>(a,b,M,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
  3294
  shows "M = [(a,c)]\<bullet>[(b,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
  3295
using a
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3296
apply(auto simp: alpha_fresh fresh_prod fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3298
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
  3299
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
  3300
9be4ab2acc13 split Class.thy into parts to 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
lemma alpha_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
  3302
  fixes M::"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
  3303
  and   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
  3304
  assumes a: "[x].M = [y].N" "z\<sharp>(x,y,M,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
  3305
  shows "M = [(x,z)]\<bullet>[(y,z)]\<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
  3306
using a
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3307
apply(auto simp: alpha_fresh fresh_prod fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  3309
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
  3310
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
  3311
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3312
lemma alpha_name_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
  3313
  fixes M::"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
  3314
  and   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
  3315
  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
  3316
  assumes a: "[x].[b].M = [y].[c].N" "z\<sharp>(x,y,M,N)" "a\<sharp>(b,c,M,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
  3317
  shows "M = [(x,z)]\<bullet>[(b,a)]\<bullet>[(c,a)]\<bullet>[(y,z)]\<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
  3318
using a
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3319
apply(auto simp: alpha_fresh fresh_prod fresh_atm 
36277
9be4ab2acc13 split Class.thy into parts to 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
                     abs_supp fin_supp abs_fresh abs_perm 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
  3321
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
  3322
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  3324
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
  3325
9be4ab2acc13 split Class.thy into parts to 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
lemma Cut_l_redu_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3327
  assumes a: "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
  3328
  shows "(\<exists>b. R = M[a\<turnstile>c>b]) \<or> (\<exists>y. R = N[x\<turnstile>n>y]) \<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
  3329
  (\<exists>y M' b N'. M = NotR (y).M' a \<and> N = NotL <b>.N' x \<and> R = Cut <b>.N' (y).M' \<and> fic M a \<and> fin N x) \<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
  3330
  (\<exists>b M1 c M2 y N'. M = AndR <b>.M1 <c>.M2 a \<and> N = AndL1 (y).N' x \<and> R = Cut <b>.M1 (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
  3331
                                                                            \<and> fic M a \<and> fin N x) \<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
  3332
  (\<exists>b M1 c M2 y N'. M = AndR <b>.M1 <c>.M2 a \<and> N = AndL2 (y).N' x \<and> R = Cut <c>.M2 (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
  3333
                                                                            \<and> fic M a \<and> fin N x) \<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
  3334
  (\<exists>b N' z M1 y M2. M = OrR1 <b>.N' a \<and> N = OrL (z).M1 (y).M2 x \<and> R = Cut <b>.N' (z).M1 
9be4ab2acc13 split Class.thy into parts to 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
                                                                            \<and> fic M a \<and> fin N x) \<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
  3336
  (\<exists>b N' z M1 y M2. M = OrR2 <b>.N' a \<and> N = OrL (z).M1 (y).M2 x \<and> R = Cut <b>.N' (y).M2 
9be4ab2acc13 split Class.thy into parts to 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
                                                                            \<and> fic M a \<and> fin N x) \<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
  3338
  (\<exists>z b M' c N1 y N2. M = ImpR (z).<b>.M' a \<and> N = ImpL <c>.N1 (y).N2 x \<and> 
9be4ab2acc13 split Class.thy into parts to 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
            R = Cut <b>.(Cut <c>.N1 (z).M') (y).N2 \<and> b\<sharp>(c,N1) \<and> fic M a \<and> 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
  3340
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
  3341
apply(erule_tac l_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
  3342
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
  3343
(* ax 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
  3344
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
  3345
apply(rule_tac x="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
  3346
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
  3347
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
  3348
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
  3349
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3351
apply(simp add: 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
  3352
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
  3353
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
  3354
(* ax 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
  3355
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
  3356
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
  3357
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
  3358
apply(thin_tac "[a].M = [aa].Ax 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
  3359
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
  3360
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
  3361
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3363
apply(simp add: 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
  3364
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
  3365
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
  3366
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
  3367
(* not 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
  3368
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
  3369
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
  3370
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
  3371
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
  3372
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
  3373
apply(drule_tac c="c" in  alpha_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
  3374
apply(simp add: fresh_prod fresh_atm 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
  3375
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
  3376
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
  3377
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
  3378
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
  3379
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
  3380
apply(simp add: calc_atm abs_fresh fresh_prod fresh_atm 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
  3381
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
  3382
apply(drule_tac z="ca" in  alpha_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
  3383
apply(simp add: fresh_prod fresh_atm 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
  3384
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
  3385
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
  3386
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
  3387
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3388
apply(auto simp: calc_atm abs_fresh fresh_left)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(case_tac "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
  3390
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
  3391
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
  3392
apply(case_tac "aa=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
  3393
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
  3394
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
  3395
(* and1 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
  3396
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
  3397
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
  3398
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
  3399
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
  3400
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
  3401
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
  3402
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
  3403
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
  3404
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
  3405
apply(drule_tac c="c" in  alpha_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
  3406
apply(simp add: fresh_prod fresh_atm 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
  3407
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3408
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
  3409
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
  3410
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
  3411
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" 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
  3412
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
  3413
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
  3414
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
  3415
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
  3416
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
  3417
apply(drule_tac z="ca" in  alpha_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
  3418
apply(simp add: fresh_prod fresh_atm 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
  3419
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  3421
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
  3422
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" 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
  3423
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
  3424
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3425
apply(auto simp: fresh_left calc_atm abs_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  3427
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
  3428
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
  3429
apply(drule_tac z="cb" in  alpha_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
  3430
apply(simp add: fresh_prod fresh_atm 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
  3431
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3433
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
  3434
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3435
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
  3436
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3437
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  3439
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
  3440
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
  3441
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
  3442
apply(drule_tac z="cb" in  alpha_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
  3443
apply(simp add: fresh_prod fresh_atm 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
  3444
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3446
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
  3447
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3448
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
  3449
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3450
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  3452
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
  3453
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
  3454
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
  3455
apply(drule_tac z="cb" in  alpha_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
  3456
apply(simp add: fresh_prod fresh_atm 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
  3457
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3458
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
  3459
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
  3460
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3461
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
  3462
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3463
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3465
(* and2 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
  3466
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
  3467
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
  3468
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
  3469
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
  3470
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
  3471
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
  3472
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
  3473
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
  3474
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
  3475
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
  3476
apply(drule_tac c="c" in  alpha_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
  3477
apply(simp add: fresh_prod fresh_atm 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
  3478
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  3480
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
  3481
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" 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
  3482
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
  3483
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
  3484
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
  3485
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
  3486
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
  3487
apply(drule_tac z="ca" in  alpha_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
  3488
apply(simp add: fresh_prod fresh_atm 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
  3489
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3491
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
  3492
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" 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
  3493
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
  3494
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3495
apply(auto simp: fresh_left calc_atm abs_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3496
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
  3497
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
  3498
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
  3499
apply(drule_tac z="cb" in  alpha_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
  3500
apply(simp add: fresh_prod fresh_atm 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
  3501
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  3503
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
  3504
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3505
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
  3506
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3507
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3509
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
  3510
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
  3511
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
  3512
apply(drule_tac z="cb" in  alpha_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
  3513
apply(simp add: fresh_prod fresh_atm 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
  3514
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  3516
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
  3517
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3518
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
  3519
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3520
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3522
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
  3523
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
  3524
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
  3525
apply(drule_tac z="cb" in  alpha_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
  3526
apply(simp add: fresh_prod fresh_atm 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
  3527
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3529
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
  3530
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3531
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
  3532
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3533
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3534
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
  3535
(* or1 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
  3536
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
  3537
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
  3538
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
  3539
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
  3540
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
  3541
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
  3542
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
  3543
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
  3544
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
  3545
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
  3546
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
  3547
apply(drule_tac c="c" in  alpha_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
  3548
apply(simp add: fresh_prod fresh_atm 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
  3549
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3551
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
  3552
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" 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
  3553
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
  3554
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
  3555
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
  3556
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
  3557
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
  3558
apply(drule_tac z="ca" in  alpha_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
  3559
apply(simp add: fresh_prod fresh_atm 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
  3560
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3561
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
  3562
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
  3563
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
  3564
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" 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
  3565
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
  3566
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3567
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  3569
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
  3570
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
  3571
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
  3572
apply(drule_tac z="cb" in  alpha_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
  3573
apply(simp add: fresh_prod fresh_atm 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
  3574
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3575
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
  3576
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
  3577
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
  3578
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3579
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
  3580
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3581
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3583
(* or2 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
  3584
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
  3585
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
  3586
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
  3587
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
  3588
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
  3589
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
  3590
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
  3591
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
  3592
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
  3593
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
  3594
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
  3595
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
  3596
apply(drule_tac c="c" in  alpha_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
  3597
apply(simp add: fresh_prod fresh_atm 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
  3598
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3599
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
  3600
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
  3601
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" 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
  3602
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
  3603
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
  3604
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
  3605
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
  3606
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
  3607
apply(drule_tac z="ca" in  alpha_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
  3608
apply(simp add: fresh_prod fresh_atm 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
  3609
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3611
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
  3612
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" 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
  3613
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
  3614
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3615
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  3617
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
  3618
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
  3619
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
  3620
apply(drule_tac z="cb" in  alpha_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
  3621
apply(simp add: fresh_prod fresh_atm 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
  3622
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3624
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
  3625
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" 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
  3626
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
  3627
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3628
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3630
(* imp-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
  3631
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
  3632
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
  3633
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
  3634
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
  3635
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
  3636
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
  3637
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
  3638
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
  3639
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
  3640
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
  3641
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
  3642
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
  3643
apply(drule_tac c="ca" in  alpha_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
  3644
apply(simp add: fresh_prod fresh_atm 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
  3645
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3647
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
  3648
apply(rule_tac s="a" and t="[(a,ca)]\<bullet>[(b,ca)]\<bullet>b" 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
  3649
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
  3650
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
  3651
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
  3652
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
  3653
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
  3654
apply(drule_tac z="cb" in  alpha_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
  3655
apply(simp add: fresh_prod fresh_atm 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
  3656
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3658
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
  3659
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(z,cb)]\<bullet>z" 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
  3660
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
  3661
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3662
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3663
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
  3664
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
  3665
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
  3666
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
  3667
apply(drule_tac z="cc" in  alpha_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
  3668
apply(simp add: fresh_prod fresh_atm 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
  3669
apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3670
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
  3671
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
  3672
apply(rule_tac s="x" and t="[(x,cc)]\<bullet>[(z,cc)]\<bullet>z" 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
  3673
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
  3674
apply(rule refl)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3675
apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3676
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
  3677
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
  3678
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3679
inductive
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3680
  c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>c _" [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3681
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3682
  left[intro]:  "\<lbrakk>\<not>fic M a; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c M{a:=(x).N}"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3683
| right[intro]: "\<lbrakk>\<not>fin N x; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c N{x:=<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
  3684
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3685
equivariance c_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
  3686
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3687
nominal_inductive c_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
  3688
 by (simp_all 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
  3689
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3690
lemma better_left[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3691
  shows "\<not>fic M a \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c M{a:=(x).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
  3692
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
  3693
  assume not_fic: "\<not>fic M 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
  3694
  obtain x'::"name" where fs1: "x'\<sharp>(N,M,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3695
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M,N)" by (rule exists_fresh(2), rule fin_supp, 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
  3696
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3697
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3698
  also have "\<dots> \<longrightarrow>\<^sub>c ([(a',a)]\<bullet>M){a':=(x').([(x',x)]\<bullet>N)}" using fs1 fs2 not_fic
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3699
    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
  3700
    apply(rule 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
  3701
    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
  3702
    apply(drule_tac a'="a" in fic_rename)
9be4ab2acc13 split Class.thy into parts to 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(simp add: perm_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
  3704
    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
  3705
    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
  3706
  also have "\<dots> = M{a:=(x).N}" using fs1 fs2
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3707
    by (simp add: subst_rename[symmetric] fresh_atm fresh_prod 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
  3708
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3710
9be4ab2acc13 split Class.thy into parts to 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
lemma better_right[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3712
  shows "\<not>fin N x \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c N{x:=<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
  3713
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
  3714
  assume not_fin: "\<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
  3715
  obtain x'::"name" where fs1: "x'\<sharp>(N,M,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3716
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M,N)" by (rule exists_fresh(2), rule fin_supp, 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
  3717
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3718
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3719
  also have "\<dots> \<longrightarrow>\<^sub>c ([(x',x)]\<bullet>N){x':=<a'>.([(a',a)]\<bullet>M)}" using fs1 fs2 not_fin
36277
9be4ab2acc13 split Class.thy into parts to 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 -
9be4ab2acc13 split Class.thy into parts to 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(rule right)
9be4ab2acc13 split Class.thy into parts to 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(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
  3723
    apply(drule_tac x'="x" in fin_rename)
9be4ab2acc13 split Class.thy into parts to 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: perm_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
  3725
    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
  3726
    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
  3727
  also have "\<dots> = N{x:=<a>.M}" using fs1 fs2
9be4ab2acc13 split Class.thy into parts to 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
    by (simp add: subst_rename[symmetric] fresh_atm fresh_prod 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
  3729
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3731
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_c_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
  3733
  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
  3734
  and   c::"coname"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3735
  shows "M \<longrightarrow>\<^sub>c M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3736
  and   "M \<longrightarrow>\<^sub>c M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>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
  3737
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
  3738
apply(induct rule: c_redu.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3739
apply(auto simp: abs_fresh rename_fresh subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(induct rule: c_redu.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3741
apply(auto simp: abs_fresh rename_fresh subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  3743
9be4ab2acc13 split Class.thy into parts to 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
inductive
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3745
  a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>a _" [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to 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
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3747
  al_redu[intro]: "M\<longrightarrow>\<^sub>l M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3748
| ac_redu[intro]: "M\<longrightarrow>\<^sub>c M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3749
| a_Cut_l: "\<lbrakk>a\<sharp>N; x\<sharp>M; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M' (x).N"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3750
| a_Cut_r: "\<lbrakk>a\<sharp>N; x\<sharp>M; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M (x).N'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3751
| a_NotL[intro]: "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^sub>a NotL <a>.M' x"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3752
| a_NotR[intro]: "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^sub>a NotR (x).M' a"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3753
| a_AndR_l: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M' <b>.N c"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3754
| a_AndR_r: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M <b>.N' c"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3755
| a_AndL1: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a AndL1 (x).M' y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3756
| a_AndL2: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a AndL2 (x).M' y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3757
| a_OrL_l: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M' (y).N z"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3758
| a_OrL_r: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M (y).N' z"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3759
| a_OrR1: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a OrR1 <a>.M' b"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3760
| a_OrR2: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a OrR2 <a>.M' b"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3761
| a_ImpL_l: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M' (x).N y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3762
| a_ImpL_r: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M (x).N' y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3763
| a_ImpR: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a ImpR (x).<a>.M' 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
  3764
9be4ab2acc13 split Class.thy into parts to 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
lemma 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
  3766
  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
  3767
  and   c::"coname"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3768
  shows "M \<longrightarrow>\<^sub>a M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3769
  and   "M \<longrightarrow>\<^sub>a M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>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
  3770
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
  3771
apply(induct rule: a_redu.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
  3772
apply(simp add: fresh_l_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
  3773
apply(simp add: fresh_c_redu)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3774
apply(auto simp: abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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(induct rule: a_redu.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
  3776
apply(simp add: fresh_l_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
  3777
apply(simp add: fresh_c_redu)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3778
apply(auto simp: abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3779
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
  3780
9be4ab2acc13 split Class.thy into parts to 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
equivariance 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
  3782
9be4ab2acc13 split Class.thy into parts to 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
nominal_inductive 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
  3784
  by (simp_all add: abs_fresh fresh_atm fresh_prod abs_supp fin_supp 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
  3785
9be4ab2acc13 split Class.thy into parts to 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
lemma better_CutL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3787
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M' (x).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
  3788
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3789
  assume red: "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
  3790
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3791
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, 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
  3792
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3793
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3794
  also have "\<dots> \<longrightarrow>\<^sub>a  Cut <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N)" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3795
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.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
  3796
  also have "\<dots> = Cut <a>.M' (x).N" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3797
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3798
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3800
9be4ab2acc13 split Class.thy into parts to 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
lemma better_CutR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3802
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M (x).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
  3803
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3804
  assume red: "N\<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
  3805
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3806
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, 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
  3807
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3808
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3809
  also have "\<dots> \<longrightarrow>\<^sub>a  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N')" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.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
  3811
  also have "\<dots> = Cut <a>.M (x).N'" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3812
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3813
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3815
    
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndRL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3817
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M' <b>.N 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
  3818
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3819
  assume red: "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
  3820
  obtain b'::"coname" where fs1: "b'\<sharp>(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, 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
  3821
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, 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
  3822
  have "AndR <a>.M <b>.N c =  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3823
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3824
  also have "\<dots> \<longrightarrow>\<^sub>a  AndR <a'>.([(a',a)]\<bullet>M') <b'>.([(b',b)]\<bullet>N) c" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3826
  also have "\<dots> = AndR <a>.M' <b>.N c" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3827
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3828
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3830
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3831
lemma better_AndRR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3832
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M <b>.N' 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
  3833
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3834
  assume red: "N\<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
  3835
  obtain b'::"coname" where fs1: "b'\<sharp>(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, 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
  3836
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, 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
  3837
  have "AndR <a>.M <b>.N c =  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3838
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3839
  also have "\<dots> \<longrightarrow>\<^sub>a  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N') c" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3840
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3841
  also have "\<dots> = AndR <a>.M <b>.N' c" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3842
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3843
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3845
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndL1_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3847
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a AndL1 (x).M' 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
  3848
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3849
  assume red: "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
  3850
  obtain x'::"name" where fs1: "x'\<sharp>(M,y,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3851
  have "AndL1 (x).M y = AndL1 (x').([(x',x)]\<bullet>M) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3852
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3853
  also have "\<dots> \<longrightarrow>\<^sub>a AndL1 (x').([(x',x)]\<bullet>M') y" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3855
  also have "\<dots> = AndL1 (x).M' y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3856
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3857
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3859
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndL2_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3861
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a AndL2 (x).M' 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
  3862
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3863
  assume red: "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
  3864
  obtain x'::"name" where fs1: "x'\<sharp>(M,y,x)" by (rule exists_fresh(1), rule fin_supp, 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
  3865
  have "AndL2 (x).M y = AndL2 (x').([(x',x)]\<bullet>M) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3866
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3867
  also have "\<dots> \<longrightarrow>\<^sub>a AndL2 (x').([(x',x)]\<bullet>M') y" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3869
  also have "\<dots> = AndL2 (x).M' y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3870
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3871
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3873
9be4ab2acc13 split Class.thy into parts to 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
lemma better_OrLL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3875
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M' (y).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
  3876
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3877
  assume red: "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
  3878
  obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, 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
  3879
  obtain y'::"name" where fs2: "y'\<sharp>(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, 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
  3880
  have "OrL (x).M (y).N z =  OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3881
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3882
  also have "\<dots> \<longrightarrow>\<^sub>a OrL (x').([(x',x)]\<bullet>M') (y').([(y',y)]\<bullet>N) z" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3884
  also have "\<dots> = OrL (x).M' (y).N z" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3885
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3886
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3888
9be4ab2acc13 split Class.thy into parts to 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
lemma better_OrLR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3890
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M (y).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
  3891
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3892
  assume red: "N\<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
  3893
  obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, 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
  3894
  obtain y'::"name" where fs2: "y'\<sharp>(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, 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
  3895
  have "OrL (x).M (y).N z =  OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3896
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3897
  also have "\<dots> \<longrightarrow>\<^sub>a OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N') z" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3899
  also have "\<dots> = OrL (x).M (y).N' z" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3900
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3901
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3903
9be4ab2acc13 split Class.thy into parts to 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
lemma better_OrR1_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3905
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a OrR1 <a>.M' 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
  3906
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3907
  assume red: "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
  3908
  obtain a'::"coname" where fs1: "a'\<sharp>(M,b,a)" by (rule exists_fresh(2), rule fin_supp, 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
  3909
  have "OrR1 <a>.M b = OrR1 <a'>.([(a',a)]\<bullet>M) b"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3910
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3911
  also have "\<dots> \<longrightarrow>\<^sub>a OrR1 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3913
  also have "\<dots> = OrR1 <a>.M' b" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3914
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3915
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3917
9be4ab2acc13 split Class.thy into parts to 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
lemma better_OrR2_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3919
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a OrR2 <a>.M' 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
  3920
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3921
  assume red: "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
  3922
  obtain a'::"coname" where fs1: "a'\<sharp>(M,b,a)" by (rule exists_fresh(2), rule fin_supp, 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
  3923
  have "OrR2 <a>.M b = OrR2 <a'>.([(a',a)]\<bullet>M) b"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3924
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3925
  also have "\<dots> \<longrightarrow>\<^sub>a OrR2 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3926
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3927
  also have "\<dots> = OrR2 <a>.M' b" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3928
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3929
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3931
9be4ab2acc13 split Class.thy into parts to 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
lemma better_ImpLL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3933
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M' (x).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
  3934
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3935
  assume red: "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
  3936
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, 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
  3937
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, 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
  3938
  have "ImpL <a>.M (x).N y =  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3939
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3940
  also have "\<dots> \<longrightarrow>\<^sub>a  ImpL <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N) y" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3942
  also have "\<dots> = ImpL <a>.M' (x).N y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3943
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3944
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3945
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
  3946
9be4ab2acc13 split Class.thy into parts to 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
lemma better_ImpLR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3948
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M (x).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
  3949
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3950
  assume red: "N\<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
  3951
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, 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
  3952
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, 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
  3953
  have "ImpL <a>.M (x).N y =  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3954
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3955
  also have "\<dots> \<longrightarrow>\<^sub>a  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N') y" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3957
  also have "\<dots> = ImpL <a>.M (x).N' y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3958
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3959
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3961
9be4ab2acc13 split Class.thy into parts to 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
lemma better_ImpR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3963
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a ImpR (x).<a>.M' 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
  3964
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3965
  assume red: "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
  3966
  obtain a'::"coname" where fs2: "a'\<sharp>(M,a,b)" by (rule exists_fresh(2), rule fin_supp, 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
  3967
  have "ImpR (x).<a>.M b = ImpR (x).<a'>.([(a',a)]\<bullet>M) b"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3968
    using fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3969
  also have "\<dots> \<longrightarrow>\<^sub>a ImpR (x).<a'>.([(a',a)]\<bullet>M') b" using fs2 red
36277
9be4ab2acc13 split Class.thy into parts to 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
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt 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
  3971
  also have "\<dots> = ImpR (x).<a>.M' b" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3972
    using fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_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
  3973
  finally show ?thesis by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  3975
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  3976
text \<open>axioms do not reduce\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma ax_do_not_l_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3979
  shows "Ax x a \<longrightarrow>\<^sub>l M \<Longrightarrow> False"
36277
9be4ab2acc13 split Class.thy into parts to 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
by (erule_tac l_redu.cases) (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
  3981
 
9be4ab2acc13 split Class.thy into parts to 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
lemma ax_do_not_c_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3983
  shows "Ax x a \<longrightarrow>\<^sub>c M \<Longrightarrow> False"
36277
9be4ab2acc13 split Class.thy into parts to 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
by (erule_tac c_redu.cases) (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
  3985
9be4ab2acc13 split Class.thy into parts to 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
lemma ax_do_not_a_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3987
  shows "Ax x a \<longrightarrow>\<^sub>a M \<Longrightarrow> False"
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3989
apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(drule ax_do_not_l_reduce)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(drule ax_do_not_c_reduce)
9be4ab2acc13 split Class.thy into parts to 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
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  3995
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_NotL_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3997
  assumes a: "NotL <a>.M x \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3998
  shows "\<exists>M'. R = NotL <a>.M' x \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  3999
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4001
apply(erule_tac l_redu.cases, 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
  4002
apply(erule_tac c_redu.cases, 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
  4003
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  4004
apply(rotate_tac 2)
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4006
apply(erule_tac l_redu.cases, 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
  4007
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4008
apply(auto simp: alpha a_redu.eqvt)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4010
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4011
apply(simp add: perm_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
  4012
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4013
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4014
apply(simp add: perm_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
  4015
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
  4016
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_NotR_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4018
  assumes a: "NotR (x).M a \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4019
  shows "\<exists>M'. R = NotR (x).M' a \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4020
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4022
apply(erule_tac l_redu.cases, 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
  4023
apply(erule_tac c_redu.cases, 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
  4024
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  4025
apply(rotate_tac 2)
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(erule_tac a_redu.cases, 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
  4027
apply(erule_tac l_redu.cases, 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
  4028
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4029
apply(auto simp: alpha a_redu.eqvt)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4031
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4032
apply(simp add: perm_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
  4033
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4034
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4035
apply(simp add: perm_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
  4036
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
  4037
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_AndR_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4039
  assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4040
  shows "(\<exists>M'. R = AndR <a>.M' <b>.N c \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = AndR <a>.M <b>.N' c \<and> N\<longrightarrow>\<^sub>aN')"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4041
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4043
apply(erule_tac l_redu.cases, 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
  4044
apply(erule_tac c_redu.cases, 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
  4045
apply(rotate_tac 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
  4046
apply(erule_tac a_redu.cases, 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
  4047
apply(erule_tac l_redu.cases, 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
  4048
apply(erule_tac c_redu.cases, 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
  4049
apply(rule disjI1)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4050
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aa)]\<bullet>M'a)" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4052
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4054
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4056
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4058
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4060
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4062
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4064
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4066
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule disjI2)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4068
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(b,ba)]\<bullet>N')" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4070
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(b,baa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4072
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(b,ba)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4074
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(b,baa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4076
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(b,ba)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4078
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(b,baa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4080
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(b,ba)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4082
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4084
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rotate_tac 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
  4086
apply(erule_tac a_redu.cases, 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
  4087
apply(erule_tac l_redu.cases, 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
  4088
apply(erule_tac c_redu.cases, 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
  4089
apply(rule disjI1)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4090
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M')" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4092
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4094
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4096
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4098
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4100
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4102
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4104
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4106
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule disjI2)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4108
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(b,ba)]\<bullet>N'a)" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4110
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(b,ba)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4112
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4114
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4116
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4117
apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4118
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4120
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4122
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4124
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4126
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_AndL1_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4128
  assumes a: "AndL1 (x).M y \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4129
  shows "\<exists>M'. R = AndL1 (x).M' y \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4130
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4132
apply(erule_tac l_redu.cases, 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
  4133
apply(erule_tac c_redu.cases, 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
  4134
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  4135
apply(rotate_tac 3)
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4137
apply(erule_tac l_redu.cases, 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
  4138
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4139
apply(auto simp: alpha a_redu.eqvt)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4141
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4142
apply(simp add: perm_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
  4143
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4144
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4145
apply(simp add: perm_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
  4146
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
  4147
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_AndL2_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4149
  assumes a: "AndL2 (x).M y \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4150
  shows "\<exists>M'. R = AndL2 (x).M' y \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4151
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4153
apply(erule_tac l_redu.cases, 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
  4154
apply(erule_tac c_redu.cases, 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
  4155
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  4156
apply(rotate_tac 3)
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4158
apply(erule_tac l_redu.cases, 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
  4159
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4160
apply(auto simp: alpha a_redu.eqvt)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4162
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4163
apply(simp add: perm_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
  4164
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4165
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4166
apply(simp add: perm_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
  4167
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
  4168
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_OrL_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4170
  assumes a: "OrL (x).M (y).N z\<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4171
  shows "(\<exists>M'. R = OrL (x).M' (y).N z \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = OrL (x).M (y).N' z \<and> N\<longrightarrow>\<^sub>aN')"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4172
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4174
apply(erule_tac l_redu.cases, 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
  4175
apply(erule_tac c_redu.cases, 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
  4176
apply(rotate_tac 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
  4177
apply(erule_tac a_redu.cases, 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
  4178
apply(erule_tac l_redu.cases, 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
  4179
apply(erule_tac c_redu.cases, 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
  4180
apply(rule disjI1)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4181
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4183
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4185
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4186
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4187
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4189
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4191
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4193
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4195
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4197
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule disjI2)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4199
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4201
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4203
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4204
apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4205
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4206
apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4207
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,ya)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4209
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4211
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,ya)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4213
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4215
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4216
apply(rotate_tac 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
  4217
apply(erule_tac a_redu.cases, 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
  4218
apply(erule_tac l_redu.cases, 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
  4219
apply(erule_tac c_redu.cases, 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
  4220
apply(rule disjI1)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4221
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M')" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4223
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4225
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4227
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4229
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4230
apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4231
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4233
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4235
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4237
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule disjI2)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4239
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4241
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4243
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4244
apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4245
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4247
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4249
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4251
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4253
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4255
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4257
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4258
lemma a_redu_OrR1_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4259
  assumes a: "OrR1 <a>.M b \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4260
  shows "\<exists>M'. R = OrR1 <a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4261
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4263
apply(erule_tac l_redu.cases, 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
  4264
apply(erule_tac c_redu.cases, 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
  4265
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  4266
apply(rotate_tac 3)
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4268
apply(erule_tac l_redu.cases, 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
  4269
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4270
apply(auto simp: alpha a_redu.eqvt)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4272
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4273
apply(simp add: perm_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
  4274
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4275
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4276
apply(simp add: perm_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
  4277
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
  4278
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4279
lemma a_redu_OrR2_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4280
  assumes a: "OrR2 <a>.M b \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4281
  shows "\<exists>M'. R = OrR2 <a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4282
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4284
apply(erule_tac l_redu.cases, 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
  4285
apply(erule_tac c_redu.cases, 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
  4286
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  4287
apply(rotate_tac 3)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4288
apply(erule_tac a_redu.cases, 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
  4289
apply(erule_tac l_redu.cases, 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
  4290
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4291
apply(auto simp: alpha a_redu.eqvt)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4293
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4294
apply(simp add: perm_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
  4295
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4296
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_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
  4297
apply(simp add: perm_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
  4298
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
  4299
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_ImpL_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4301
  assumes a: "ImpL <a>.M (y).N z\<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4302
  shows "(\<exists>M'. R = ImpL <a>.M' (y).N z \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = ImpL <a>.M (y).N' z \<and> N\<longrightarrow>\<^sub>aN')"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4303
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4305
apply(erule_tac l_redu.cases, 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
  4306
apply(erule_tac c_redu.cases, 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
  4307
apply(rotate_tac 5)
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4309
apply(erule_tac l_redu.cases, 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
  4310
apply(erule_tac c_redu.cases, 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
  4311
apply(rule disjI1)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4312
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4314
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4316
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4318
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4319
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4320
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4322
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4324
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4326
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4328
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule disjI2)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4330
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N')" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4332
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(y,xa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4334
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4336
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4338
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4339
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4340
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(y,xa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4342
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(y,xa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4344
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(y,xa)]\<bullet>N')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4346
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rotate_tac 5)
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4349
apply(erule_tac l_redu.cases, 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
  4350
apply(erule_tac c_redu.cases, 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
  4351
apply(rule disjI1)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4352
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M')" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4354
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4356
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4358
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4360
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4362
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4364
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4366
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4368
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule disjI2)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4370
apply(auto simp: alpha a_redu.eqvt)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(y,xa)]\<bullet>N'a)" in exI) 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4372
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4374
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4376
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4378
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4380
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4382
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4383
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4384
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4386
apply(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4388
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_ImpR_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4390
  assumes a: "ImpR (x).<a>.M b \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4391
  shows "\<exists>M'. R = ImpR (x).<a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  4392
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4394
apply(erule_tac l_redu.cases, 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
  4395
apply(erule_tac c_redu.cases, 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
  4396
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  4397
apply(rotate_tac 3)
36277
9be4ab2acc13 split Class.thy into parts to 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(erule_tac a_redu.cases, 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
  4399
apply(erule_tac l_redu.cases, 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
  4400
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4401
apply(auto simp: alpha a_redu.eqvt abs_perm abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4402
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4403
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4405
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4407
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4409
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4411
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4413
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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_tac x="([(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4415
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  4417
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
  4418
apply(rule perm_compose)
9be4ab2acc13 split Class.thy into parts to 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
apply(simp add: calc_atm perm_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
  4420
apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4421
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  4423
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
  4424
apply(rule perm_compose)
9be4ab2acc13 split Class.thy into parts to 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(simp add: calc_atm perm_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
  4426
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4427
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4429
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule_tac x="([(a,aa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4431
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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 perm_compose)
9be4ab2acc13 split Class.thy into parts to 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 add: calc_atm perm_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
  4436
apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4437
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4439
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
  4440
apply(rule perm_compose)
9be4ab2acc13 split Class.thy into parts to 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 add: calc_atm perm_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
  4442
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
  4443
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  4444
text \<open>Transitive Closure\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
abbreviation
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4447
 a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>a* _" [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to 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
where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 66453
diff changeset
  4449
  "M \<longrightarrow>\<^sub>a* M' \<equiv> (a_redu)\<^sup>*\<^sup>* M 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
  4450
9be4ab2acc13 split Class.thy into parts to 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
lemma a_starI:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4452
  assumes a: "M \<longrightarrow>\<^sub>a M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4453
  shows "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
  4454
using a 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
  4455
9be4ab2acc13 split Class.thy into parts to 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
lemma a_starE:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4457
  assumes a: "M \<longrightarrow>\<^sub>a* M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4458
  shows "M = M' \<or> (\<exists>N. M \<longrightarrow>\<^sub>a N \<and> N \<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
  4459
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
  4460
by (induct) (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
  4461
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_refl:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4463
  shows "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
  4464
  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
  4465
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_trans[trans]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4467
  assumes a1: "M1\<longrightarrow>\<^sub>a* M2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4468
  and     a2: "M2\<longrightarrow>\<^sub>a* M3"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4469
  shows "M1 \<longrightarrow>\<^sub>a* M3"
36277
9be4ab2acc13 split Class.thy into parts to 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
using a2 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
  4471
by (induct) (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
  4472
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  4473
text \<open>congruence rules for \<open>\<longrightarrow>\<^sub>a*\<close>\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma ax_do_not_a_star_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4476
  shows "Ax x a \<longrightarrow>\<^sub>a* M \<Longrightarrow> M = Ax x 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
  4477
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4478
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
  4479
apply(drule  ax_do_not_a_reduce)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  4482
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_CutL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4485
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M' (x).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
  4486
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_CutR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4489
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M (x).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
  4490
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_Cut:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4493
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M' (x).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
  4494
by (blast intro!: a_star_CutL a_star_CutR intro: rtranclp_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
  4495
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_NotR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4497
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^sub>a* NotR (x).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
  4498
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_NotL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4501
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^sub>a* NotL <a>.M' 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
  4502
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndRL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4505
    "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M' <b>.N 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
  4506
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndRR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4509
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M <b>.N' 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
  4510
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4513
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M' <b>.N' 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
  4514
by (blast intro!: a_star_AndRL a_star_AndRR intro: rtranclp_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
  4515
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4516
lemma a_star_AndL1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4517
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a* AndL1 (x).M' 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
  4518
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndL2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4521
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a* AndL2 (x).M' 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
  4522
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrLL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4525
    "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M' (y).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
  4526
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrLR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4529
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M (y).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
  4530
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4533
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M' (y).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
  4534
by (blast intro!: a_star_OrLL a_star_OrLR intro: rtranclp_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
  4535
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrR1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4537
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a* OrR1 <a>.M' 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
  4538
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrR2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4541
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a* OrR2 <a>.M' 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
  4542
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_ImpLL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4545
    "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M' (y).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
  4546
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_ImpLR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4549
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M (y).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
  4550
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_ImpL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4553
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M' (y).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
  4554
by (blast intro!: a_star_ImpLL a_star_ImpLR intro: rtranclp_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
  4555
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_ImpR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4557
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a* ImpR (x).<a>.M' 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
  4558
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemmas a_star_congs = a_star_Cut a_star_NotR a_star_NotL a_star_AndR a_star_AndL1 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
  4561
                      a_star_OrL a_star_OrR1 a_star_OrR2 a_star_ImpL 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
  4562
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_NotL_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4564
  assumes a: "NotL <a>.M x \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4565
  shows "\<exists>M'. R = NotL <a>.M' x \<and> 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
  4566
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
  4567
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4569
apply(drule a_redu_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
  4570
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
  4571
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
  4572
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_NotR_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4574
  assumes a: "NotR (x).M a \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4575
  shows "\<exists>M'. R = NotR (x).M' a \<and> 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
  4576
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
  4577
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4578
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
  4579
apply(drule a_redu_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
  4580
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
  4581
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
  4582
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_AndR_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4584
  assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4585
  shows "(\<exists>M' N'. R = AndR <a>.M' <b>.N' c \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<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
  4586
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
  4587
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4589
apply(drule a_redu_AndR_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4590
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4592
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_AndL1_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4594
  assumes a: "AndL1 (x).M y \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4595
  shows "\<exists>M'. R = AndL1 (x).M' y \<and> 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
  4596
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
  4597
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4599
apply(drule a_redu_AndL1_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4600
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4602
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_AndL2_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4604
  assumes a: "AndL2 (x).M y \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4605
  shows "\<exists>M'. R = AndL2 (x).M' y \<and> 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
  4606
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
  4607
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4609
apply(drule a_redu_AndL2_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4610
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4612
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4613
lemma a_star_redu_OrL_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4614
  assumes a: "OrL (x).M (y).N z \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4615
  shows "(\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<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
  4616
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
  4617
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4619
apply(drule a_redu_OrL_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4620
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4622
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_OrR1_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4624
  assumes a: "OrR1 <a>.M y \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4625
  shows "\<exists>M'. R = OrR1 <a>.M' y \<and> 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
  4626
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
  4627
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4629
apply(drule a_redu_OrR1_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4630
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4632
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_OrR2_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4634
  assumes a: "OrR2 <a>.M y \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4635
  shows "\<exists>M'. R = OrR2 <a>.M' y \<and> 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
  4636
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
  4637
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4639
apply(drule a_redu_OrR2_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4640
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4642
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_ImpR_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4644
  assumes a: "ImpR (x).<a>.M y \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4645
  shows "\<exists>M'. R = ImpR (x).<a>.M' y \<and> 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
  4646
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
  4647
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4649
apply(drule a_redu_ImpR_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4650
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4652
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_ImpL_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4654
  assumes a: "ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4655
  shows "(\<exists>M' N'. R = ImpL <a>.M' (y).N' z \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<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
  4656
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
  4657
apply(induct set: rtranclp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4659
apply(drule a_redu_ImpL_elim)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4660
apply(auto simp: alpha trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4662
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  4663
text \<open>Substitution\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_not_fin1:
9be4ab2acc13 split Class.thy into parts to 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
  shows "\<not>fin(M{x:=<c>.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
  4667
apply(nominal_induct M avoiding: x c P 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
  4668
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
  4669
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(erule fin.cases, 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
  4672
apply(erule fin.cases, 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
  4673
apply(erule fin.cases, 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
  4674
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},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
  4675
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
  4676
apply(simp add: 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
  4677
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
  4678
apply(simp add: fresh_fun_simp_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
  4679
apply(erule fin.cases, 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
  4680
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(erule fin.cases, 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
  4682
apply(erule fin.cases, 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
  4683
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,name1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  4685
apply(simp add: 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
  4686
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
  4687
apply(simp add: fresh_fun_simp_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
  4688
apply(erule fin.cases, 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
  4689
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(erule fin.cases, 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
  4691
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,name1)")
9be4ab2acc13 split Class.thy into parts to 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
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
  4693
apply(simp add: 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
  4694
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
  4695
apply(simp add: fresh_fun_simp_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
  4696
apply(erule fin.cases, 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
  4697
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(erule fin.cases, 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
  4699
apply(erule fin.cases, 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
  4700
apply(erule fin.cases, 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
  4701
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<c>.P},P,name1,trm2{x:=<c>.P},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
  4702
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
  4703
apply(simp add: 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
  4704
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
  4705
apply(simp add: fresh_fun_simp_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
  4706
apply(erule fin.cases, 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
  4707
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(erule fin.cases, 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
  4709
apply(erule fin.cases, 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
  4710
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<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
  4711
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
  4712
apply(simp add: 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
  4713
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
  4714
apply(simp add: fresh_fun_simp_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
  4715
apply(erule fin.cases, 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
  4716
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(erule fin.cases, 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
  4718
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
  4719
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_not_fin2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "\<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
  4722
  shows "\<not>fin(M{c:=(x).P}) 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
  4723
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
  4724
apply(nominal_induct M avoiding: x c P 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
  4725
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
  4726
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname:=(x).P},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
  4730
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
  4731
apply(simp add: 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
  4732
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
  4733
apply(simp add: fresh_fun_simp_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
  4734
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule freshn_after_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
  4739
apply(simp add: 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
  4740
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm1{coname3:=(x).P},P,coname1,trm2{coname3:=(x).P},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
  4741
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
  4742
apply(simp add: 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
  4743
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
  4744
apply(simp add: fresh_fun_simp_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
  4745
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4749
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
  4750
apply(drule freshn_after_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
  4751
apply(simp add: fin.intros 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
  4752
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4754
apply(drule freshn_after_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
  4755
apply(simp add: fin.intros 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
  4756
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(x).P},P,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  4758
apply(simp add: 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
  4759
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
  4760
apply(simp add: fresh_fun_simp_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
  4761
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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 exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(x).P},P,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  4766
apply(simp add: 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
  4767
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
  4768
apply(simp add: fresh_fun_simp_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
  4769
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4774
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
  4775
apply(drule freshn_after_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
  4776
apply(drule freshn_after_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
  4777
apply(simp add: fin.intros 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
  4778
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(x).P},P,coname1,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
  4779
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
  4780
apply(simp add: 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
  4781
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
  4782
apply(simp add: fresh_fun_simp_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
  4783
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4785
apply(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fin_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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)[1]
9be4ab2acc13 split Class.thy into parts to 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 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
  4789
apply(drule freshn_after_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
  4790
apply(drule freshn_after_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
  4791
apply(simp add: fin.intros 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
  4792
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
  4793
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_not_fic1:
9be4ab2acc13 split Class.thy into parts to 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
  shows "\<not>fic (M{a:=(x).P}) 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
  4796
apply(nominal_induct M avoiding: a x P 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
  4797
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
  4798
apply(erule fic.cases, 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
  4799
apply(erule fic.cases, 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
  4800
apply(erule fic.cases, 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
  4801
apply(erule fic.cases, 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
  4802
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname:=(x).P},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
  4803
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
  4804
apply(simp add: 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
  4805
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
  4806
apply(simp add: fresh_fun_simp_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
  4807
apply(erule fic.cases, 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
  4808
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(erule fic.cases, 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
  4810
apply(erule fic.cases, 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
  4811
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,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
  4812
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
  4813
apply(simp add: 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
  4814
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
  4815
apply(simp add: fresh_fun_simp_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
  4816
apply(erule fic.cases, 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
  4817
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(erule fic.cases, 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
  4819
apply(erule fic.cases, 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
  4820
apply(erule fic.cases, 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
  4821
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(x).P},P,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  4823
apply(simp add: 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
  4824
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
  4825
apply(simp add: fresh_fun_simp_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
  4826
apply(erule fic.cases, 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
  4827
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(erule fic.cases, 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
  4829
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(x).P},P,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  4831
apply(simp add: 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
  4832
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
  4833
apply(simp add: fresh_fun_simp_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
  4834
apply(erule fic.cases, 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
  4835
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
apply(erule fic.cases, 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
  4837
apply(erule fic.cases, 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
  4838
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(x).P},P,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  4840
apply(simp add: 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
  4841
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
  4842
apply(simp add: fresh_fun_simp_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
  4843
apply(erule fic.cases, 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
  4844
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(erule fic.cases, 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
  4846
apply(erule fic.cases, 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
  4847
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
  4848
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_not_fic2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "\<not>fic M 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
  4851
  shows "\<not>fic(M{x:=<b>.P}) 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
  4852
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
  4853
apply(nominal_induct M avoiding: x a P 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
  4854
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
  4855
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule freshc_after_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
  4860
apply(simp add: 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
  4861
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.P},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
  4862
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
  4863
apply(simp add: 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
  4864
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
  4865
apply(simp add: fresh_fun_simp_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
  4866
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4869
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4870
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
  4871
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
  4872
apply(drule freshc_after_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
  4873
apply(drule freshc_after_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
  4874
apply(simp add: fic.intros 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
  4875
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.P},P,name1)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4876
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
  4877
apply(simp add: 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
  4878
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
  4879
apply(simp add: fresh_fun_simp_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
  4880
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4882
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.P},P,name1)")
9be4ab2acc13 split Class.thy into parts to 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
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
  4885
apply(simp add: 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
  4886
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
  4887
apply(simp add: fresh_fun_simp_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
  4888
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  4893
apply(drule freshc_after_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
  4894
apply(simp add: fic.intros 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
  4895
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  4897
apply(drule freshc_after_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
  4898
apply(simp add: fic.intros 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
  4899
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.P},P,name1,trm2{x:=<b>.P},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
  4900
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
  4901
apply(simp add: 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
  4902
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
  4903
apply(simp add: fresh_fun_simp_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
  4904
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  4909
apply(drule freshc_after_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
  4910
apply(simp add: fic.intros 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
  4911
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.P},trm2{name2:=<b>.P},P,name1)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4912
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
  4913
apply(simp add: 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
  4914
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
  4915
apply(simp add: fresh_fun_simp_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
  4916
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
apply(drule fic_elims, simp)
9be4ab2acc13 split Class.thy into parts to 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
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
  4920
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  4921
text \<open>Reductions\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4922
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_l_reduce:
9be4ab2acc13 split Class.thy into parts to 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
  assumes  a: "fin M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4925
  and      b: "M \<longrightarrow>\<^sub>l 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
  4926
  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
  4927
using 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
  4928
apply(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
  4929
apply(erule fin.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
  4930
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
  4931
apply(rotate_tac 3)
9be4ab2acc13 split Class.thy into parts to 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
apply(erule fin.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
  4933
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
  4934
apply(erule fin.cases, 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
  4935
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
  4936
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_c_reduce:
9be4ab2acc13 split Class.thy into parts to 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
  assumes  a: "fin M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4939
  and      b: "M \<longrightarrow>\<^sub>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
  4940
  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
  4941
using 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
  4942
apply(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
  4943
apply(erule fin.cases, 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
  4944
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
  4945
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_a_reduce:
9be4ab2acc13 split Class.thy into parts to 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
  assumes  a: "fin M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4948
  and      b: "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
  4949
  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
  4950
using 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
  4951
apply(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
  4952
apply(drule ax_do_not_a_reduce)
9be4ab2acc13 split Class.thy into parts to 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
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule a_redu_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
  4955
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
  4956
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
  4957
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
  4958
apply(drule a_redu_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
  4959
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
  4960
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
  4961
apply(force simp add: abs_fresh 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
  4962
apply(drule a_redu_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
  4963
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
  4964
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
  4965
apply(force simp add: abs_fresh 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
  4966
apply(drule a_redu_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
  4967
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
  4968
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
  4969
apply(force simp add: abs_fresh 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
  4970
apply(force simp add: abs_fresh 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
  4971
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
  4972
apply(force simp add: abs_fresh 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
  4973
apply(force simp add: abs_fresh 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
  4974
apply(drule a_redu_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
  4975
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
  4976
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
  4977
apply(force simp add: abs_fresh 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
  4978
apply(force simp add: abs_fresh 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
  4979
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
  4980
apply(force simp add: abs_fresh 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
  4981
apply(force simp add: abs_fresh 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
  4982
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
  4983
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_a_star_reduce:
9be4ab2acc13 split Class.thy into parts to 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
  assumes  a: "fin M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4986
  and      b: "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
  4987
  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
  4988
using 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
  4989
apply(induct set: rtranclp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4990
apply(auto simp: fin_a_reduce)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4992
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_l_reduce:
9be4ab2acc13 split Class.thy into parts to 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
  assumes  a: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4995
  and      b: "M \<longrightarrow>\<^sub>l 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
  4996
  shows "fic 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
  4997
using 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
  4998
apply(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
  4999
apply(erule fic.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
  5000
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
  5001
apply(rotate_tac 3)
9be4ab2acc13 split Class.thy into parts to 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
apply(erule fic.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
  5003
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
  5004
apply(erule fic.cases, 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
  5005
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
  5006
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_c_reduce:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5009
  and     b: "M \<longrightarrow>\<^sub>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
  5010
  shows "fic 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
  5011
using 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
  5012
apply(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
  5013
apply(erule fic.cases, 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
  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 fic_a_reduce:
9be4ab2acc13 split Class.thy into parts to 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: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5018
  and     b: "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
  5019
  shows "fic 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
  5020
using 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
  5021
apply(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
  5022
apply(drule ax_do_not_a_reduce)
9be4ab2acc13 split Class.thy into parts to 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
apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
apply(drule a_redu_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
  5025
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
  5026
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
  5027
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
  5028
apply(drule a_redu_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
  5029
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
  5030
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
  5031
apply(force simp add: abs_fresh 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
  5032
apply(force simp add: abs_fresh 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
  5033
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
  5034
apply(force simp add: abs_fresh 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
  5035
apply(force simp add: abs_fresh 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
  5036
apply(drule a_redu_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
  5037
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
  5038
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
  5039
apply(force simp add: abs_fresh 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
  5040
apply(drule a_redu_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
  5041
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
  5042
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
  5043
apply(force simp add: abs_fresh 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
  5044
apply(drule a_redu_ImpR_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
  5045
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
  5046
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
  5047
apply(force simp add: abs_fresh 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
  5048
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
  5049
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_a_star_reduce:
9be4ab2acc13 split Class.thy into parts to 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
  assumes  a: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5052
  and      b: "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
  5053
  shows "fic 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
  5054
using 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
  5055
apply(induct set: rtranclp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5056
apply(auto simp: fic_a_reduce)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  5058
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  5059
text \<open>substitution properties\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_with_ax1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5062
  shows "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[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
  5063
proof(nominal_induct M avoiding: x a 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
  5064
  case (Ax z b x a y)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5065
  show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[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
  5066
  proof (cases "z=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
  5067
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "z=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
  5069
    have "(Ax z b){x:=<a>.Ax y a} = Cut <a>.Ax y a (x).Ax x b" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5070
    also have "\<dots> \<longrightarrow>\<^sub>a* (Ax x b)[x\<turnstile>n>y]" by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5071
    finally show "Ax z b{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]" using eq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5074
    assume neq: "z\<noteq>x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5075
    then show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]" using neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5077
next
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut b M z N x a 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
  5079
  have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>N" "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>M" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5080
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5081
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5082
  show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[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
  5083
  proof (cases "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
  5084
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "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
  5086
    have "(Cut <b>.M (z).N){x:=<a>.Ax y a} = Cut <a>.Ax y a (z).(N{x:=<a>.Ax y a})" using fs eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5087
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.Ax y a (z).(N[x\<turnstile>n>y])" using ih2 a_star_congs by blast
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using eq
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5089
      by (simp add: trm.inject alpha calc_atm fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5090
    finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5093
    assume neq: "M \<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
  5094
    have "(Cut <b>.M (z).N){x:=<a>.Ax y a} = Cut <b>.(M{x:=<a>.Ax y a}) (z).(N{x:=<a>.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
  5095
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5096
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using ih1 ih2 a_star_congs by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5097
    finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5098
  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
  5099
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotR z M b x a 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
  5101
  have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>b" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5102
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(NotR (z).M b){x:=<a>.Ax y a} = NotR (z).(M{x:=<a>.Ax y a}) b" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5104
  also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[x\<turnstile>n>y]) b" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5105
  finally show "(NotR (z).M b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotR (z).M b)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5106
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL b M z x a 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
  5108
  have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5109
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5110
  show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[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
  5111
  proof(cases "z=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
  5112
    case True
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5113
    assume eq: "z=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
  5114
    obtain x'::"name" where new: "x'\<sharp>(Ax y a,M{x:=<a>.Ax y a})" by (rule exists_fresh(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5115
    have "(NotL <b>.M z){x:=<a>.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
  5116
                        fresh_fun (\<lambda>x'. Cut <a>.Ax y a (x').NotL <b>.(M{x:=<a>.Ax y a}) 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
  5117
      using eq fs by simp
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5118
    also have "\<dots> = Cut <a>.Ax y a (x').NotL <b>.(M{x:=<a>.Ax y a}) 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
  5119
      using new by (simp add: fresh_fun_simp_NotL fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5120
    also have "\<dots> \<longrightarrow>\<^sub>a* (NotL <b>.(M{x:=<a>.Ax y a}) x')[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
  5121
      using new 
9be4ab2acc13 split Class.thy into parts to 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(rule_tac 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
  5123
      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
  5124
      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
  5125
      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
  5126
      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
  5127
    also have "\<dots> = NotL <b>.(M{x:=<a>.Ax y a}) y" using new by (simp add: nrename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5128
    also have "\<dots> \<longrightarrow>\<^sub>a* NotL <b>.(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (NotL <b>.M z)[x\<turnstile>n>y]" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5130
    finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5133
    assume neq: "z\<noteq>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
  5134
    have "(NotL <b>.M z){x:=<a>.Ax y a} = NotL <b>.(M{x:=<a>.Ax y a}) z" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5135
    also have "\<dots> \<longrightarrow>\<^sub>a* NotL <b>.(M[x\<turnstile>n>y]) z" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5136
    finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]" using neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5138
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR c M d N e x a 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
  5140
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "d\<sharp>x" "d\<sharp>a" "d\<sharp>y" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5141
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5142
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} = AndR <c>.(M{x:=<a>.Ax y a}) <d>.(N{x:=<a>.Ax y a}) 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
  5144
    using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5145
  also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[x\<turnstile>n>y]) <d>.(N[x\<turnstile>n>y]) e" using ih1 ih2 by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5146
  finally show "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[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
  5147
    using fs by simp
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5149
  case (AndL1 u M v x a 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
  5150
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5151
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5152
  show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[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
  5153
  proof(cases "v=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
  5154
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "v=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
  5156
    obtain v'::"name" where new: "v'\<sharp>(Ax y a,M{x:=<a>.Ax y a},u)" by (rule exists_fresh(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    have "(AndL1 (u).M v){x:=<a>.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
  5158
                        fresh_fun (\<lambda>v'. Cut <a>.Ax y a (v').AndL1 (u).(M{x:=<a>.Ax y a}) 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
  5159
      using eq fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a>.Ax y a (v').AndL1 (u).(M{x:=<a>.Ax y a}) 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
  5161
      using new by (simp add: fresh_fun_simp_AndL1 fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5162
    also have "\<dots> \<longrightarrow>\<^sub>a* (AndL1 (u).(M{x:=<a>.Ax y a}) v')[v'\<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
  5163
      using new 
9be4ab2acc13 split Class.thy into parts to 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_tac 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
  5165
      apply(rule a_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
  5166
      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
  5167
      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
  5168
      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
  5169
      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
  5170
    also have "\<dots> = AndL1 (u).(M{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5171
      by (auto simp: fresh_prod fresh_atm nrename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5172
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (AndL1 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5174
    finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5177
    assume neq: "v\<noteq>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
  5178
    have "(AndL1 (u).M v){x:=<a>.Ax y a} = AndL1 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5179
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5180
    finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5182
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5183
  case (AndL2 u M v x a 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
  5184
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5185
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5186
  show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[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
  5187
  proof(cases "v=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
  5188
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "v=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
  5190
    obtain v'::"name" where new: "v'\<sharp>(Ax y a,M{x:=<a>.Ax y a},u)" by (rule exists_fresh(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    have "(AndL2 (u).M v){x:=<a>.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
  5192
                        fresh_fun (\<lambda>v'. Cut <a>.Ax y a (v').AndL2 (u).(M{x:=<a>.Ax y a}) 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
  5193
      using eq fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a>.Ax y a (v').AndL2 (u).(M{x:=<a>.Ax y a}) 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
  5195
      using new by (simp add: fresh_fun_simp_AndL2 fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5196
    also have "\<dots> \<longrightarrow>\<^sub>a* (AndL2 (u).(M{x:=<a>.Ax y a}) v')[v'\<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
  5197
      using new 
9be4ab2acc13 split Class.thy into parts to 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_tac 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
  5199
      apply(rule a_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
  5200
      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
  5201
      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
  5202
      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
  5203
      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
  5204
    also have "\<dots> = AndL2 (u).(M{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5205
      by (auto simp: fresh_prod fresh_atm nrename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5206
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (AndL2 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5208
    finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5211
    assume neq: "v\<noteq>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
  5212
    have "(AndL2 (u).M v){x:=<a>.Ax y a} = AndL2 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5213
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5214
    finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5216
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR1 c M d  x a 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
  5218
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5219
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(OrR1 <c>.M d){x:=<a>.Ax y a} = OrR1 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5221
  also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5222
  finally show "(OrR1 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR2 c M d  x a 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
  5225
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5226
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(OrR2 <c>.M d){x:=<a>.Ax y a} = OrR2 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5228
  also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5229
  finally show "(OrR2 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL u M v N z x a 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
  5232
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "v\<sharp>x" "v\<sharp>a" "v\<sharp>y" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5233
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5234
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5235
  show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[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
  5236
  proof(cases "z=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
  5237
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "z=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
  5239
    obtain z'::"name" where new: "z'\<sharp>(Ax y a,M{x:=<a>.Ax y a},N{x:=<a>.Ax y a},u,v,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
  5240
      by (rule exists_fresh(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    have "(OrL (u).M (v).N z){x:=<a>.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
  5242
                 fresh_fun (\<lambda>z'. Cut <a>.Ax y a (z').OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) 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
  5243
      using eq fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a>.Ax y a (z').OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) 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
  5245
      using new by (simp add: fresh_fun_simp_OrL)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5246
    also have "\<dots> \<longrightarrow>\<^sub>a* (OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z')[z'\<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
  5247
      using new 
9be4ab2acc13 split Class.thy into parts to 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_tac 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
  5249
      apply(rule a_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
  5250
      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
  5251
      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
  5252
      apply(simp_all 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
  5253
      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
  5254
    also have "\<dots> = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5255
      by (auto simp: fresh_prod fresh_atm nrename_fresh subst_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5256
    also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) 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
  5257
      using ih1 ih2 by (auto intro: a_star_congs)
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (OrL (u).M (v).N z)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5259
    finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5262
    assume neq: "z\<noteq>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
  5263
    have "(OrL (u).M (v).N z){x:=<a>.Ax y a} = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) 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
  5264
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5265
    also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) 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
  5266
      using ih1 ih2 by (auto intro: a_star_congs)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5267
    finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5269
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpR z c M d x a 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
  5271
  have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "z\<sharp>d" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5272
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(ImpR (z).<c>.M d){x:=<a>.Ax y a} = ImpR (z).<c>.(M{x:=<a>.Ax y a}) d" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5274
  also have "\<dots> \<longrightarrow>\<^sub>a* ImpR (z).<c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5275
  finally show "(ImpR (z).<c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL c M u N v x a 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
  5278
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5279
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5280
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5281
  show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[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
  5282
  proof(cases "v=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
  5283
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "v=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
  5285
    obtain v'::"name" where new: "v'\<sharp>(Ax y a,M{x:=<a>.Ax y a},N{x:=<a>.Ax y a},y,a,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
  5286
      by (rule exists_fresh(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5287
    have "(ImpL <c>.M (u).N v){x:=<a>.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
  5288
                 fresh_fun (\<lambda>v'. Cut <a>.Ax y a (v').ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) 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
  5289
      using eq fs by simp 
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a>.Ax y a (v').ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) 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
  5291
      using new by (simp add: fresh_fun_simp_ImpL)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5292
    also have "\<dots> \<longrightarrow>\<^sub>a* (ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v')[v'\<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
  5293
      using new 
9be4ab2acc13 split Class.thy into parts to 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(rule_tac 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
  5295
      apply(rule a_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
  5296
      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
  5297
      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
  5298
      apply(simp_all 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
  5299
      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
  5300
    also have "\<dots> = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5301
      by (auto simp: fresh_prod subst_fresh fresh_atm trm.inject alpha rename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5302
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) 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
  5303
      using ih1 ih2 by (auto intro: a_star_congs)
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5305
    finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5306
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5308
    assume neq: "v\<noteq>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
  5309
    have "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) 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
  5310
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5311
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) v" 
36277
9be4ab2acc13 split Class.thy into parts to 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
      using ih1 ih2 by (auto intro: a_star_congs)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5313
    finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[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
  5314
      using fs neq by simp
9be4ab2acc13 split Class.thy into parts to 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
  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
  5316
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
  5317
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_with_ax2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5319
  shows "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<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
  5320
proof(nominal_induct M avoiding: b a x 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
  5321
  case (Ax z c b a x)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5322
  show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<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
  5323
  proof (cases "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
  5324
    case True
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5325
    assume eq: "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
  5326
    have "(Ax z c){b:=(x).Ax x a} = Cut <b>.Ax z c (x).Ax x a" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5327
    also have "\<dots> \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" using eq by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5328
    finally show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5331
    assume neq: "c\<noteq>b"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5332
    then show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5334
next
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut c M z N b a 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
  5336
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>N" "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>M" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5337
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5338
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5339
  show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<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
  5340
  proof (cases "N = 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
  5341
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "N = 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
  5343
    have "(Cut <c>.M (z).N){b:=(x).Ax x a} = Cut <c>.(M{b:=(x).Ax x a}) (x).Ax x a" using eq fs by simp 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5344
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <c>.(M[b\<turnstile>c>a]) (x).Ax x a" using ih1 a_star_congs by blast
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using eq fs
9be4ab2acc13 split Class.thy into parts to 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
      by (simp add: trm.inject alpha calc_atm fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5347
    finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5350
    assume neq: "N \<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
  5351
    have "(Cut <c>.M (z).N){b:=(x).Ax x a} = Cut <c>.(M{b:=(x).Ax x a}) (z).(N{b:=(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
  5352
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5353
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using ih1 ih2 a_star_congs by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5354
    finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5356
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5357
  case (NotR z M c b a 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
  5358
  have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>c" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5359
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5360
  show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<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
  5361
  proof (cases "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
  5362
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "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
  5364
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a})" by (rule exists_fresh(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    have "(NotR (z).M c){b:=(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
  5366
                        fresh_fun (\<lambda>a'. Cut <a'>.NotR (z).M{b:=(x).Ax x a} a' (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
  5367
      using eq fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a'>.NotR (z).M{b:=(x).Ax x a} a' (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
  5369
      using new by (simp add: fresh_fun_simp_NotR fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5370
    also have "\<dots> \<longrightarrow>\<^sub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\<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
  5371
      using new 
9be4ab2acc13 split Class.thy into parts to 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_tac 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
  5373
      apply(rule a_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
  5374
      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
  5375
      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
  5376
      apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5377
      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
  5378
    also have "\<dots> = NotR (z).(M{b:=(x).Ax x a}) a" using new by (simp add: crename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5379
    also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (NotR (z).M c)[b\<turnstile>c>a]" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5381
    finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5382
  next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5383
    case 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
  5384
    assume neq: "c\<noteq>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
  5385
    have "(NotR (z).M c){b:=(x).Ax x a} = NotR (z).(M{b:=(x).Ax x a}) c" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5386
    also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[b\<turnstile>c>a]) c" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5387
    finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]" using neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5389
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL c M z b a 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
  5391
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5392
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(NotL <c>.M z){b:=(x).Ax x a} = NotL <c>.(M{b:=(x).Ax x a}) z" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5394
  also have "\<dots> \<longrightarrow>\<^sub>a* NotL <c>.(M[b\<turnstile>c>a]) z" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5395
  finally show "(NotL <c>.M z){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotL <c>.M z)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR c M d N e b a 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
  5398
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "d\<sharp>b" "d\<sharp>a" "d\<sharp>x" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5399
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5400
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5401
  show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<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
  5402
  proof(cases "e=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
  5403
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "e=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
  5405
    obtain e'::"coname" where new: "e'\<sharp>(Ax x a,M{b:=(x).Ax x a},N{b:=(x).Ax x a},c,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
  5406
      by (rule exists_fresh(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    have "(AndR <c>.M <d>.N e){b:=(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
  5408
               fresh_fun (\<lambda>e'. Cut <e'>.AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e' (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
  5409
      using eq fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <e'>.AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e' (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
  5411
      using new by (simp add: fresh_fun_simp_AndR fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5412
    also have "\<dots> \<longrightarrow>\<^sub>a* (AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e')[e'\<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
  5413
      using new 
9be4ab2acc13 split Class.thy into parts to 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(rule_tac 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
  5415
      apply(rule a_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
  5416
      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
  5417
      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
  5418
      apply(simp_all 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
  5419
      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
  5420
    also have "\<dots> = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5421
      by (auto simp: fresh_prod fresh_atm subst_fresh crename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5422
    also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) a" using ih1 ih2 by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5424
    finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5427
    assume neq: "e\<noteq>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
  5428
    have "(AndR <c>.M <d>.N e){b:=(x).Ax x a} = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) 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
  5429
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5430
    also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) e" using ih1 ih2 by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5431
    finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<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
  5432
      using fs neq by simp
9be4ab2acc13 split Class.thy into parts to 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
  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
  5434
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 u M v b a 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
  5436
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5437
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(AndL1 (u).M v){b:=(x).Ax x a} = AndL1 (u).(M{b:=(x).Ax x a}) v" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5439
  also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5440
  finally show "(AndL1 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5441
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5442
  case (AndL2 u M v b a 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
  5443
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5444
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(AndL2 (u).M v){b:=(x).Ax x a} = AndL2 (u).(M{b:=(x).Ax x a}) v" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5446
  also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5447
  finally show "(AndL2 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR1 c M d  b a 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
  5450
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5451
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5452
  show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<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
  5453
  proof(cases "d=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
  5454
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "d=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
  5456
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a},c,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
  5457
      by (rule exists_fresh(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    have "(OrR1 <c>.M d){b:=(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
  5459
             fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp)
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a'>.OrR1 <c>.M{b:=(x).Ax x a} a' (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
  5461
      using new by (simp add: fresh_fun_simp_OrR1)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5462
    also have "\<dots> \<longrightarrow>\<^sub>a* (OrR1 <c>.M{b:=(x).Ax x a} a')[a'\<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
  5463
      using new 
9be4ab2acc13 split Class.thy into parts to 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_tac 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
  5465
      apply(rule a_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
  5466
      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
  5467
      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
  5468
      apply(simp_all 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
  5469
      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
  5470
    also have "\<dots> = OrR1 <c>.M{b:=(x).Ax x a} a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5471
      by (auto simp: fresh_prod fresh_atm crename_fresh subst_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5472
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (OrR1 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5474
    finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5477
    assume neq: "d\<noteq>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
  5478
    have "(OrR1 <c>.M d){b:=(x).Ax x a} = OrR1 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5479
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5480
    finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5482
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5483
  case (OrR2 c M d  b a 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
  5484
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5485
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5486
  show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<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
  5487
  proof(cases "d=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
  5488
    case True
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5489
    assume eq: "d=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
  5490
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a},c,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
  5491
      by (rule exists_fresh(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    have "(OrR2 <c>.M d){b:=(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
  5493
             fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp)
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a'>.OrR2 <c>.M{b:=(x).Ax x a} a' (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
  5495
      using new by (simp add: fresh_fun_simp_OrR2)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5496
    also have "\<dots> \<longrightarrow>\<^sub>a* (OrR2 <c>.M{b:=(x).Ax x a} a')[a'\<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
  5497
      using new 
9be4ab2acc13 split Class.thy into parts to 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(rule_tac 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
  5499
      apply(rule a_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
  5500
      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
  5501
      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
  5502
      apply(simp_all 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
  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
    also have "\<dots> = OrR2 <c>.M{b:=(x).Ax x a} a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5505
      by (auto simp: fresh_prod fresh_atm crename_fresh subst_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5506
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (OrR2 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5508
    finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5511
    assume neq: "d\<noteq>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
  5512
    have "(OrR2 <c>.M d){b:=(x).Ax x a} = OrR2 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5513
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5514
    finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5516
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL u M v N z b a 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
  5518
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "v\<sharp>b" "v\<sharp>a" "v\<sharp>x" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5519
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5520
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(OrL (u).M (v).N z){b:=(x).Ax x a} = OrL (u).(M{b:=(x).Ax x a}) (v).(N{b:=(x).Ax x a}) 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
  5522
    using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5523
  also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[b\<turnstile>c>a]) (v).(N[b\<turnstile>c>a]) z" using ih1 ih2 by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5524
  finally show "(OrL (u).M (v).N z){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpR z c M d b a 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
  5527
  have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "z\<sharp>d" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5528
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5529
  show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<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
  5530
  proof(cases "b=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
  5531
    case True
9be4ab2acc13 split Class.thy into parts to 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
    assume eq: "b=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
  5533
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a},x,a,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
  5534
      by (rule exists_fresh(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5535
    have "(ImpR (z).<c>.M d){b:=(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
  5536
                fresh_fun (\<lambda>a'. Cut <a'>.ImpR z.<c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a'>.ImpR z.<c>.M{b:=(x).Ax x a} a' (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
  5538
      using new by (simp add: fresh_fun_simp_ImpR)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5539
    also have "\<dots> \<longrightarrow>\<^sub>a* (ImpR z.<c>.M{b:=(x).Ax x a} a')[a'\<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
  5540
      using new 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5541
      apply(rule_tac 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
  5542
      apply(rule a_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
  5543
      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
  5544
      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
  5545
      apply(simp_all 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
  5546
      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
  5547
    also have "\<dots> = ImpR z.<c>.M{b:=(x).Ax x a} a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5548
      by (auto simp: fresh_prod crename_fresh subst_fresh fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5549
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpR z.<c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (ImpR z.<c>.M b)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5551
    finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using eq by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case 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
  5554
    assume neq: "b\<noteq>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
  5555
    have "(ImpR (z).<c>.M d){b:=(x).Ax x a} = ImpR (z).<c>.(M{b:=(x).Ax x a}) d" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5556
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpR (z).<c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5557
    finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using neq fs by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  5559
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL c M u N v b a 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
  5561
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5562
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5563
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to 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
  have "(ImpL <c>.M (u).N v){b:=(x).Ax x a} = ImpL <c>.(M{b:=(x).Ax x a}) (u).(N{b:=(x).Ax x a}) 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
  5565
    using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5566
  also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[b\<turnstile>c>a]) (u).(N[b\<turnstile>c>a]) v" 
36277
9be4ab2acc13 split Class.thy into parts to 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
    using ih1 ih2 by (auto intro: a_star_congs)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  5568
  finally show "(ImpL <c>.M (u).N v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[b\<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
  5569
    using fs by simp
9be4ab2acc13 split Class.thy into parts to 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
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
  5571
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  5572
text \<open>substitution lemmas\<close>
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5574
lemma not_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
  5575
  shows "\<not>(b\<sharp>M) \<Longrightarrow> M{b:=(y).Q} \<noteq> 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
  5576
apply(nominal_induct M avoiding: b y Q x a rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5577
apply(auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(y).Q},Q)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5580
apply(simp add: 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
  5581
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
  5582
apply(simp add: fresh_fun_simp_NotR 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
  5583
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(y).Q},Q)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5586
apply(simp add: 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
  5587
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
  5588
apply(simp add: fresh_fun_simp_NotR 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
  5589
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,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
  5591
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
  5592
apply(simp add: 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
  5593
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
  5594
apply(simp add: fresh_fun_simp_AndR 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
  5595
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,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
  5597
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
  5598
apply(simp add: 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
  5599
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
  5600
apply(simp add: fresh_fun_simp_AndR 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
  5601
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,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
  5603
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
  5604
apply(simp add: 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
  5605
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
  5606
apply(simp add: fresh_fun_simp_AndR 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
  5607
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5609
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
  5610
apply(simp add: 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
  5611
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
  5612
apply(simp add: fresh_fun_simp_OrR1 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
  5613
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
9be4ab2acc13 split Class.thy into parts to 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 add: 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
  5617
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
  5618
apply(simp add: fresh_fun_simp_OrR1 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
  5619
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5622
apply(simp add: 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
  5623
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
  5624
apply(simp add: fresh_fun_simp_OrR2 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
  5625
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5628
apply(simp add: 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
  5629
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
  5630
apply(simp add: fresh_fun_simp_OrR2 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
  5631
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5634
apply(simp add: 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
  5635
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
  5636
apply(simp add: fresh_fun_simp_ImpR 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
  5637
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5640
apply(simp add: 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
  5641
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
  5642
apply(simp add: fresh_fun_simp_ImpR 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
  5643
apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
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
  5645
9be4ab2acc13 split Class.thy into parts to 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
lemma not_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
  5647
  shows "\<not>(x\<sharp>M) \<Longrightarrow> M{x:=<b>.Q} \<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
  5648
apply(nominal_induct M avoiding: b y Q x a rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5649
apply(auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5652
apply(simp add: 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
  5653
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
  5654
apply(simp add: fresh_fun_simp_NotL 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
  5655
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5658
apply(simp add: 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
  5659
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
  5660
apply(simp add: fresh_fun_simp_NotL 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
  5661
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5664
apply(simp add: 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
  5665
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
  5666
apply(simp add: fresh_fun_simp_AndL1 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
  5667
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5669
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
  5670
apply(simp add: 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
  5671
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
  5672
apply(simp add: fresh_fun_simp_AndL1 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
  5673
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5675
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
  5676
apply(simp add: 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
  5677
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
  5678
apply(simp add: fresh_fun_simp_AndL2 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
  5679
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5682
apply(simp add: 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
  5683
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
  5684
apply(simp add: fresh_fun_simp_AndL2 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
  5685
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.Q},Q,trm2{x:=<b>.Q},name1,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
  5687
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
  5688
apply(simp add: 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
  5689
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
  5690
apply(simp add: fresh_fun_simp_OrL 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
  5691
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.Q},Q,trm2{x:=<b>.Q},name1,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
  5693
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
  5694
apply(simp add: 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
  5695
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
  5696
apply(simp add: fresh_fun_simp_OrL 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
  5697
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.Q},Q,trm2{x:=<b>.Q},name1,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
  5699
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
  5700
apply(simp add: 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
  5701
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
  5702
apply(simp add: fresh_fun_simp_OrL 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
  5703
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.Q},Q,trm2{name2:=<b>.Q},name1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5706
apply(simp add: 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
  5707
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
  5708
apply(simp add: fresh_fun_simp_ImpL 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
  5709
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.Q},Q,trm2{name2:=<b>.Q},name1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5712
apply(simp add: 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
  5713
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
  5714
apply(simp add: fresh_fun_simp_ImpL 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
  5715
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.Q},Q,trm2{name2:=<b>.Q},name1)")
9be4ab2acc13 split Class.thy into parts to 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(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
  5718
apply(simp add: 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
  5719
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
  5720
apply(simp add: fresh_fun_simp_ImpL 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
  5721
apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
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
  5723
9be4ab2acc13 split Class.thy into parts to 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
lemma interesting_subst1:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5725
  assumes a: "x\<noteq>y" "x\<sharp>P" "y\<sharp>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
  5726
  shows "N{y:=<c>.P}{x:=<c>.P} = N{x:=<c>.Ax y c}{y:=<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
  5727
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
  5728
proof(nominal_induct N avoiding: x y c P 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
  5729
  case 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
  5730
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5731
    by (auto simp: abs_fresh fresh_atm forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
next 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5733
  case (Cut d M u M' x' y' c P)
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  5734
  with assms show ?case
36277
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5737
    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
  5738
    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
  5739
    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
  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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5743
    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
  5744
    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
  5745
    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
  5746
    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
  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(simp add: trm.inject alpha 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
  5750
    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
  5751
    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
  5752
    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
  5753
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5756
    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
  5757
    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
  5758
    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
  5759
    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
  5760
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5761
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5762
    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
  5763
    apply(simp add: trm.inject alpha 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
  5764
    apply(simp add: trm.inject alpha 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
  5765
    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
  5766
    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
  5767
    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
  5768
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5771
    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
  5772
    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
  5773
    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
  5774
    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
  5775
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5778
    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
  5779
    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
  5780
    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
  5781
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  5784
    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
  5785
    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
  5786
    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
  5787
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5790
    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
  5791
    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
  5792
    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
  5793
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5796
    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
  5797
    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
  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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5801
    apply(case_tac "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
  5802
    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
  5803
    apply(simp add: not_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
  5804
    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
  5805
    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
  5806
    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
  5807
    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
  5808
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5811
    apply(case_tac "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
  5812
    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
  5813
    apply(simp add: not_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
  5814
    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
  5815
next
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  5817
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5818
    by (auto simp: abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL d M 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
  5821
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5822
    apply (auto simp: abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},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
  5824
    apply(erule exE, simp add: 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
  5825
    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
  5826
    apply(simp add: fresh_fun_simp_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
  5827
    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
  5828
    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
  5829
    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
  5830
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5832
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  5834
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},Ax y c,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
  5836
    apply(erule exE, simp only: 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
  5837
    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
  5838
    apply(simp only: fresh_fun_simp_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
  5839
    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
  5840
    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
  5841
    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
  5842
    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
  5843
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha forget 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
  5846
    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
  5847
    apply(rule substn.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
  5848
    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
  5849
    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
  5850
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  5852
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5853
  case (AndR d1 M d2 M' d3)
9be4ab2acc13 split Class.thy into parts to 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
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5855
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 u M 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
  5858
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5859
    apply(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},u,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
  5861
    apply(erule exE, simp add: 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
  5862
    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
  5863
    apply(simp add: fresh_fun_simp_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
  5864
    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
  5865
    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
  5866
    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
  5867
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5869
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  5871
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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 "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},u,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
  5873
    apply(erule exE, simp only: 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
  5874
    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
  5875
    apply(simp only: fresh_fun_simp_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
  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 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
  5878
    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
  5879
    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
  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(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5882
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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 exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  5885
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL2 u M 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
  5887
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5888
    apply(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},u,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
  5890
    apply(erule exE, simp add: 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
  5891
    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
  5892
    apply(simp add: fresh_fun_simp_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
  5893
    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
  5894
    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
  5895
    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
  5896
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5898
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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 add: trm.inject alpha 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
  5900
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},u,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
  5902
    apply(erule exE, simp only: 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
  5903
    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
  5904
    apply(simp only: fresh_fun_simp_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
  5905
    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
  5906
    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
  5907
    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
  5908
    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
  5909
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5911
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  5914
next
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  5916
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5917
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  5920
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5921
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL x1 M x2 M' 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
  5924
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5925
    apply(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<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
  5927
                                        M'{y:=<c>.P},M'{x:=<c>.Ax y c}{y:=<c>.P},x1,x2,x3,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
  5928
    apply(erule exE, simp add: 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
  5929
    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
  5930
    apply(simp add: fresh_fun_simp_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
  5931
    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
  5932
    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
  5933
    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
  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(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5936
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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 add: trm.inject alpha 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
  5938
    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
  5939
    apply(rule substn.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
  5940
    apply(simp add: abs_fresh subst_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
  5941
    apply(simp add: abs_fresh subst_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
  5942
    apply(force)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<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
  5946
                                        M'{x:=<c>.Ax y c},M'{x:=<c>.Ax y c}{y:=<c>.P},x1,x2,x3,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
  5947
    apply(erule exE, simp only: 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
  5948
    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
  5949
    apply(simp only: fresh_fun_simp_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
  5950
    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
  5951
    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
  5952
    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
  5953
    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
  5954
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  5957
    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
  5958
    apply(rule substn.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
  5959
    apply(simp add: abs_fresh subst_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
  5960
    apply(simp add: abs_fresh subst_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
  5961
    apply(force)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5962
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5963
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  5964
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  5966
next
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  5968
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5969
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL a M x1 M' x2)
9be4ab2acc13 split Class.thy into parts to 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
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5973
    apply(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x2:=<c>.P},M{x:=<c>.Ax x2 c}{x2:=<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
  5975
                                        M'{x2:=<c>.P},M'{x:=<c>.Ax x2 c}{x2:=<c>.P},x1,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
  5976
    apply(erule exE, simp add: 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
  5977
    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
  5978
    apply(simp add: fresh_fun_simp_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
  5979
    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
  5980
    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
  5981
    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
  5982
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  5984
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(simp add: trm.inject alpha 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
  5986
    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
  5987
    apply(rule substn.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
  5988
    apply(simp add: abs_fresh subst_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
  5989
    apply(simp add: abs_fresh subst_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
  5990
    apply(force)
9be4ab2acc13 split Class.thy into parts to 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 exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x2:=<c>.Ax y c},M{x2:=<c>.Ax y c}{y:=<c>.P},
48567
3c4d7ff75f01 tuned proofs -- avoid odd situations of polymorphic Frees in goal state;
wenzelm
parents: 41893
diff changeset
  5993
                                        M'{x2:=<c>.Ax y c},M'{x2:=<c>.Ax y c}{y:=<c>.P},x1,x2,y,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
  5994
    apply(erule exE, simp only: 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
  5995
    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
  5996
    apply(simp only: fresh_fun_simp_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
  5997
    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
  5998
    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
  5999
    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
  6000
    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
  6001
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
    apply(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
  6004
    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
  6005
    apply(rule substn.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
  6006
    apply(simp add: abs_fresh subst_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
  6007
    apply(simp add: abs_fresh subst_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
  6008
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6009
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6012
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
  6013
9be4ab2acc13 split Class.thy into parts to 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
lemma interesting_subst1':
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "x\<noteq>y" "x\<sharp>P" "y\<sharp>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
  6016
  shows "N{y:=<c>.P}{x:=<c>.P} = N{x:=<a>.Ax y a}{y:=<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
  6017
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
  6018
  show ?thesis
9be4ab2acc13 split Class.thy into parts to 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
  proof (cases "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
  6020
    case True then show ?thesis using a by (simp add: interesting_subst1)
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case False then show ?thesis 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
  6023
      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
  6024
      apply(subgoal_tac "N{x:=<a>.Ax y a} = N{x:=<c>.([(c,a)]\<bullet>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
  6025
      apply(simp add: interesting_subst1 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
  6026
      apply(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(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
  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
  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
  6030
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
  6031
9be4ab2acc13 split Class.thy into parts to 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
lemma interesting_subst2:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "a\<noteq>b" "a\<sharp>P" "b\<sharp>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
  6034
  shows "N{a:=(y).P}{b:=(y).P} = N{b:=(y).Ax y a}{a:=(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
  6035
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
  6036
proof(nominal_induct N avoiding: a b y P 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
  6037
  case 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
  6038
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6039
    by (auto simp: abs_fresh fresh_atm forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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
next 
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut d M u M' x' y' c P)
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6042
  with assms show ?case
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6044
    apply(auto simp: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  6046
    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
  6047
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  6049
    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
  6050
    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
  6051
    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
  6052
    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
  6053
    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
  6054
    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
  6055
    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
  6056
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  6058
    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
  6059
    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
  6060
    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
  6061
    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
  6062
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  6064
    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
  6065
    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
  6066
    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
  6067
    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
  6068
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  6070
    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
  6071
    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
  6072
    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
  6073
    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
  6074
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  6076
    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
  6077
    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
  6078
    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
  6079
    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
  6080
    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
  6081
    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
  6082
    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
  6083
    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
  6084
    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
  6085
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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
    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
  6087
    apply(simp add: fresh_atm trm.inject alpha 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
  6088
    apply(case_tac "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
  6089
    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
  6090
    apply(simp add: not_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
  6091
    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
  6092
    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
  6093
    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
  6094
    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
  6095
    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
  6096
    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
  6097
    apply(case_tac "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
  6098
    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
  6099
    apply(simp add: not_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
  6100
    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
  6101
next
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  6103
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6104
    by (auto simp: abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotR u M 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
  6107
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6108
    apply (auto simp: abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P},u,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
  6110
    apply(erule exE, simp add: 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
  6111
    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
  6112
    apply(simp add: fresh_fun_simp_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
  6113
    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
  6114
    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
  6115
    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
  6116
    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
  6117
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6118
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  6120
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P},Ax y a,y,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
  6122
    apply(erule exE, simp only: 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
  6123
    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
  6124
    apply(simp only: fresh_fun_simp_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
  6125
    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
  6126
    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
  6127
    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
  6128
    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
  6129
    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
  6130
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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 add: trm.inject alpha forget 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
  6132
    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
  6133
    apply(rule substc.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
  6134
    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
  6135
    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
  6136
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6138
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR d1 M d2 M' d3)
9be4ab2acc13 split Class.thy into parts to 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
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6141
    apply(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,M{d3:=(y).P},M{b:=(y).Ax y d3}{d3:=(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
  6143
                                        M'{d3:=(y).P},M'{b:=(y).Ax y d3}{d3:=(y).P},d1,d2,d3,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
  6144
    apply(erule exE, simp add: 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
  6145
    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
  6146
    apply(simp add: fresh_fun_simp_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
  6147
    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
  6148
    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
  6149
    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
  6150
    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
  6151
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6152
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  6154
    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
  6155
    apply(rule substc.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
  6156
    apply(simp add: abs_fresh subst_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
  6157
    apply(simp add: abs_fresh subst_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
  6158
    apply(force)
9be4ab2acc13 split Class.thy into parts to 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(simp)
9be4ab2acc13 split Class.thy into parts to 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 exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,Ax y a,M{d3:=(y).Ax y a},M{d3:=(y).Ax y a}{a:=(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
  6162
                                        M'{d3:=(y).Ax y a},M'{d3:=(y).Ax y a}{a:=(y).P},d1,d2,d3,y,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
  6163
    apply(erule exE, simp only: 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
  6164
    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
  6165
    apply(simp only: fresh_fun_simp_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
  6166
    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
  6167
    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
  6168
    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
  6169
    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
  6170
    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
  6171
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  6173
    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
  6174
    apply(rule substc.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
  6175
    apply(simp add: abs_fresh subst_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
  6176
    apply(simp add: abs_fresh subst_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
  6177
    apply(force)
9be4ab2acc13 split Class.thy into parts to 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(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6179
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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 exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6182
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 u M 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
  6184
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6185
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL2 u M 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
  6188
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6189
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR1 d 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
  6192
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6193
    apply (auto simp: abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,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
  6195
    apply(erule exE, simp add: 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
  6196
    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
  6197
    apply(simp add: fresh_fun_simp_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
  6198
    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
  6199
    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
  6200
    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
  6201
    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
  6202
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6203
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  6205
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,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
  6207
    apply(erule exE, simp only: 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
  6208
    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
  6209
    apply(simp only: fresh_fun_simp_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
  6210
    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
  6211
    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
  6212
    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
  6213
    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
  6214
    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
  6215
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha forget 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
  6217
    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
  6218
    apply(rule substc.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
  6219
    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
  6220
    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
  6221
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6223
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR2 d 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
  6225
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6226
    apply (auto simp: abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,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
  6228
    apply(erule exE, simp add: 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
  6229
    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
  6230
    apply(simp add: fresh_fun_simp_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
  6231
    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
  6232
    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
  6233
    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
  6234
    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
  6235
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6236
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  6238
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,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
  6240
    apply(erule exE, simp only: 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
  6241
    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
  6242
    apply(simp only: fresh_fun_simp_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
  6243
    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
  6244
    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
  6245
    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
  6246
    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
  6247
    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
  6248
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha forget 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
  6250
    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
  6251
    apply(rule substc.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
  6252
    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
  6253
    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
  6254
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6256
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL x1 M x2 M' 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
  6258
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6259
    by(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case 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
  6262
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6263
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpR u e M 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
  6266
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6267
    apply(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,e,d,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(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
  6269
    apply(erule exE, simp add: 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
  6270
    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
  6271
    apply(simp add: fresh_fun_simp_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
  6272
    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
  6273
    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
  6274
    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
  6275
    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
  6276
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6277
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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(simp add: trm.inject alpha 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
  6279
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(e,d,P,Ax y a,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(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
  6281
    apply(erule exE, simp only: 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
  6282
    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
  6283
    apply(simp only: fresh_fun_simp_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
  6284
    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
  6285
    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
  6286
    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
  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(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
  6289
    apply(simp)
9be4ab2acc13 split Class.thy into parts to 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(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
  6291
    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
  6292
    apply(rule substc.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
  6293
    apply(simp add: abs_fresh subst_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
  6294
    apply(simp add: abs_fresh subst_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
  6295
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6296
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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 exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6299
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
  6300
9be4ab2acc13 split Class.thy into parts to 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
lemma interesting_subst2':
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "a\<noteq>b" "a\<sharp>P" "b\<sharp>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
  6303
  shows "N{a:=(y).P}{b:=(y).P} = N{b:=(z).Ax z a}{a:=(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
  6304
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
  6305
  show ?thesis
9be4ab2acc13 split Class.thy into parts to 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
  proof (cases "z=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
  6307
    case True then show ?thesis using a by (simp add: interesting_subst2)
9be4ab2acc13 split Class.thy into parts to 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
  next
9be4ab2acc13 split Class.thy into parts to 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
    case False then show ?thesis 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
  6310
      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
  6311
      apply(subgoal_tac "N{b:=(z).Ax z a} = N{b:=(y).([(y,z)]\<bullet>Ax z 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
  6312
      apply(simp add: interesting_subst2 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
  6313
      apply(rule subst_rename)
9be4ab2acc13 split Class.thy into parts to 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(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
  6315
      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
  6316
  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
  6317
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
  6318
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_subst1:
9be4ab2acc13 split Class.thy into parts to 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
  assumes a: "a\<sharp>(Q,b)" "x\<sharp>(y,P,Q)" "b\<sharp>Q" "y\<sharp>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
  6321
  shows "M{x:=<a>.P}{b:=(y).Q} = M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}"
9be4ab2acc13 split Class.thy into parts to 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
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
  6323
proof(nominal_induct M avoiding: x a P b y Q 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
  6324
  case (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
  6325
  have fs: "a\<sharp>(Q,b)" "x\<sharp>(y,P,Q)" "b\<sharp>Q" "y\<sharp>P" by fact+
9be4ab2acc13 split Class.thy into parts to 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
  { assume asm: "z=x \<and> 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
  6327
    have "(Ax x b){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (x).Ax x b){b:=(y).Q}" using fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).Q"
9be4ab2acc13 split Class.thy into parts to 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
      using fs by (simp_all 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
  6330
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).(Q{x:=<a>.(P{b:=(y).Q})})" using fs by (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
  6331
    also have "\<dots> = (Cut <b>.Ax x b (y).Q){x:=<a>.(P{b:=(y).Q})}"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6332
      using fs asm by (auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (Ax x b){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    finally have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" 
9be4ab2acc13 split Class.thy into parts to 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
      using asm by simp
9be4ab2acc13 split Class.thy into parts to 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
  }
9be4ab2acc13 split Class.thy into parts to 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
  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
  6338
  { assume asm: "z\<noteq>x \<and> 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
  6339
    have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}" using asm by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <b>.Ax z c (y).Q" using fs asm by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <b>.(Ax z c{x:=<a>.(P{b:=(y).Q})}) (y).(Q{x:=<a>.(P{b:=(y).Q})})" 
9be4ab2acc13 split Class.thy into parts to 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
      using fs asm by (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
  6343
    also have "\<dots> = (Cut <b>.Ax z c (y).Q){x:=<a>.(P{b:=(y).Q})}" using asm fs
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6344
      by (auto simp: trm.inject subst_fresh fresh_prod fresh_atm abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using asm fs by simp
9be4ab2acc13 split Class.thy into parts to 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
    finally have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" by simp
9be4ab2acc13 split Class.thy into parts to 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
  }
9be4ab2acc13 split Class.thy into parts to 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
  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
  6349
  { assume asm: "z=x \<and> c\<noteq>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
  6350
    have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (x).Ax z c){b:=(y).Q}" using fs asm by simp
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6351
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (x).Ax z c" using fs asm by (auto simp: trm.inject abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (Ax z c){x:=<a>.(P{b:=(y).Q})}" using fs asm by simp
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using asm by 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
  6354
    finally have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" by simp
9be4ab2acc13 split Class.thy into parts to 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
  }
9be4ab2acc13 split Class.thy into parts to 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
  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
  6357
  { assume asm: "z\<noteq>x \<and> c\<noteq>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
  6358
    have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using asm by 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
  6359
  }
9be4ab2acc13 split Class.thy into parts to 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
  ultimately show ?case 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
  6361
next
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut c M z 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
  6363
  { assume asm: "M = Ax x c \<and> N = 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
  6364
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (z).(N{x:=<a>.P})){b:=(y).Q}" 
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6365
      using Cut asm by simp
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6366
    also have "\<dots> = (Cut <a>.P (z).N){b:=(y).Q}" using Cut asm by (simp add: fresh_atm)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6367
    also have "\<dots> = (Cut <a>.(P{b:=(y).Q}) (y).Q)" using Cut asm by (auto simp: fresh_prod fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.(P{b:=(y).Q}) (y).Q)" by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = (Cut <c>.M (y).Q){x:=<a>.(P{b:=(y).Q})}"
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6370
      using Cut asm by (simp add: fresh_atm)
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6371
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).(Q{x:=<a>.(P{b:=(y).Q})})" using Cut asm
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6372
      by (auto simp: fresh_prod fresh_atm subst_fresh)
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6373
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).Q" using Cut asm by (simp add: forget)
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = Cut <a>.(P{b:=(y).Q}) (y).Q"
9be4ab2acc13 split Class.thy into parts to 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
      by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" 
9be4ab2acc13 split Class.thy into parts to 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
      using eq1 eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  }
9be4ab2acc13 split Class.thy into parts to 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
  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
  6380
  { assume asm: "M \<noteq> Ax x c \<and> N = 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
  6381
    have neq: "M{b:=(y).Q} \<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
  6382
    proof (cases "b\<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
  6383
      case True then show ?thesis using asm by (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
  6384
    next
9be4ab2acc13 split Class.thy into parts to 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
      case False then show ?thesis by (simp add: not_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
  6386
    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
  6387
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.(M{x:=<a>.P}) (z).(N{x:=<a>.P})){b:=(y).Q}"
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6388
      using Cut asm by simp
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6389
    also have "\<dots> = (Cut <c>.(M{x:=<a>.P}) (z).N){b:=(y).Q}" using Cut asm by (simp add: fresh_atm)
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6390
    also have "\<dots> = Cut <c>.(M{x:=<a>.P}{b:=(y).Q}) (y).Q" using Cut asm by (simp add: abs_fresh)
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6391
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}}) (y).Q" using Cut asm by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally 
9be4ab2acc13 split Class.thy into parts to 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
    have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = Cut <c>.(M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}}) (y).Q" 
9be4ab2acc13 split Class.thy into parts to 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
      by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6396
               (Cut <c>.(M{b:=(y).Q}) (y).Q){x:=<a>.(P{b:=(y).Q})}" using Cut asm by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (y).(Q{x:=<a>.(P{b:=(y).Q})})"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6398
      using Cut asm neq by (auto simp: fresh_prod fresh_atm subst_fresh abs_fresh)
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6399
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (y).Q" using Cut asm by (simp add: forget)
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} 
9be4ab2acc13 split Class.thy into parts to 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
                                       = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (y).Q" by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" 
9be4ab2acc13 split Class.thy into parts to 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
      using eq1 eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  }
9be4ab2acc13 split Class.thy into parts to 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
  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
  6406
  { assume asm: "M = Ax x c \<and> N \<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
  6407
    have neq: "N{x:=<a>.P} \<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
  6408
    proof (cases "x\<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
  6409
      case True then show ?thesis using asm by (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
  6410
    next
9be4ab2acc13 split Class.thy into parts to 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
      case False then show ?thesis by (simp add: not_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
  6412
    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
  6413
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (z).(N{x:=<a>.P})){b:=(y).Q}"
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6414
      using Cut asm by simp
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6415
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (z).(N{x:=<a>.P}{b:=(y).Q})" using Cut asm neq 
36277
9be4ab2acc13 split Class.thy into parts to 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
      by (simp add: abs_fresh)
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6417
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" using Cut asm by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} 
9be4ab2acc13 split Class.thy into parts to 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
                    = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} 
9be4ab2acc13 split Class.thy into parts to 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
                    = (Cut <c>.(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}"
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6422
      using Cut asm by auto
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = (Cut <c>.M (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6424
      using Cut asm by (auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" 
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6426
      using Cut asm by (simp add: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally 
9be4ab2acc13 split Class.thy into parts to 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
    have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} 
9be4ab2acc13 split Class.thy into parts to 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
         = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}"
9be4ab2acc13 split Class.thy into parts to 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
      using eq1 eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  }
9be4ab2acc13 split Class.thy into parts to 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
  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
  6434
  { assume asm: "M \<noteq> Ax x c \<and> N \<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
  6435
    have neq1: "N{x:=<a>.P} \<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
  6436
    proof (cases "x\<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
  6437
      case True then show ?thesis using asm by (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
  6438
    next
9be4ab2acc13 split Class.thy into parts to 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
      case False then show ?thesis by (simp add: not_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
  6440
    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
  6441
    have neq2: "M{b:=(y).Q} \<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
  6442
    proof (cases "b\<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
  6443
      case True then show ?thesis using asm by (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
  6444
    next
9be4ab2acc13 split Class.thy into parts to 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
      case False then show ?thesis by (simp add: not_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
  6446
    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
  6447
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.(M{x:=<a>.P}) (z).(N{x:=<a>.P})){b:=(y).Q}"
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6448
      using Cut asm by simp
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6449
    also have "\<dots> = Cut <c>.(M{x:=<a>.P}{b:=(y).Q}) (z).(N{x:=<a>.P}{b:=(y).Q})" using Cut asm neq1
36277
9be4ab2acc13 split Class.thy into parts to 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
      by (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
  6451
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})"
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6452
      using Cut asm by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q}
9be4ab2acc13 split Class.thy into parts to 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
             = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6456
                (Cut <c>.(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}" using Cut asm neq1 by simp
36277
9be4ab2acc13 split Class.thy into parts to 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
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})"
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  6458
      using Cut asm neq2 by (simp add: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    finally have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
9be4ab2acc13 split Class.thy into parts to 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
           Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
9be4ab2acc13 split Class.thy into parts to 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
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}"
9be4ab2acc13 split Class.thy into parts to 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
      using eq1 eq2 by simp
9be4ab2acc13 split Class.thy into parts to 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
  }
9be4ab2acc13 split Class.thy into parts to 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
  ultimately show ?case 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
  6465
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotR z 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
  6467
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6468
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>a'::coname. a'\<sharp>(M{c:=(y).Q},M{c:=(y).Q}{x:=<a>.P{c:=(y).Q}},Q,a,P,c,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
  6470
    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
  6471
    apply(simp add: 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
  6472
    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
  6473
    apply(simp add: fresh_fun_simp_NotR 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
  6474
    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
  6475
    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
  6476
    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
  6477
    apply(simp add: fresh_prod fresh_atm subst_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
  6478
    apply(simp add: fresh_prod fresh_atm subst_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
  6479
    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
  6480
    apply(simp add: fresh_prod fresh_atm 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
  6481
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6483
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL c 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
  6485
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6486
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},P{b:=(y).Q},M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},y,Q)")
9be4ab2acc13 split Class.thy into parts to 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
    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
  6489
    apply(simp add: 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
  6490
    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
  6491
    apply(simp add: fresh_fun_simp_NotL 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
  6492
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to 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
    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
  6494
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR c1 M c2 N c3)
9be4ab2acc13 split Class.thy into parts to 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
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6497
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to 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
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{c3:=(y).Q},M{c3:=(y).Q}{x:=<a>.P{c3:=(y).Q}},c2,c3,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
  6499
                                     P{c3:=(y).Q},N{c3:=(y).Q},N{c3:=(y).Q}{x:=<a>.P{c3:=(y).Q}},c1)")
9be4ab2acc13 split Class.thy into parts to 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
    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
  6501
    apply(simp add: 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
  6502
    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
  6503
    apply(simp add: fresh_fun_simp_AndR 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
  6504
    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
  6505
    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
  6506
    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
  6507
    apply(simp_all add: fresh_atm 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
  6508
    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
  6509
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6510
    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
  6511
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6512
  case (AndL1 z1 M z2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6513
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6514
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6515
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}})")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6516
    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
  6517
    apply(simp add: 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
  6518
    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
  6519
    apply(simp add: fresh_fun_simp_AndL1 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
  6520
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6521
    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
  6522
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6523
  case (AndL2 z1 M z2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6524
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6525
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6526
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}})")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6527
    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
  6528
    apply(simp add: 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
  6529
    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
  6530
    apply(simp add: fresh_fun_simp_AndL2 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
  6531
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6532
    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
  6533
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6534
  case (OrL z1 M z2 N z3)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6535
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6536
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6537
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},z2,z3,a,y,Q,
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6538
                                     P{b:=(y).Q},N{x:=<a>.P},N{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},z1)")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6539
    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
  6540
    apply(simp add: 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
  6541
    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
  6542
    apply(simp add: fresh_fun_simp_OrL 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
  6543
    apply(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
  6544
    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
  6545
    apply(rule substc.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
  6546
    apply(simp_all add: fresh_atm 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
  6547
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6548
    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
  6549
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6550
  case (OrR1 c1 M c2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6551
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6552
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6553
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{c2:=(y).Q},a,P{c2:=(y).Q},c1,
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6554
                                                     M{c2:=(y).Q}{x:=<a>.P{c2:=(y).Q}})")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6555
    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
  6556
    apply(simp add: 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
  6557
    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
  6558
    apply(simp add: fresh_fun_simp_OrR1 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
  6559
    apply(simp_all add: fresh_atm subst_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
  6560
    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
  6561
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6562
    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
  6563
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6564
  case (OrR2 c1 M c2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6565
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6566
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6567
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{c2:=(y).Q},a,P{c2:=(y).Q},c1,
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6568
                                                     M{c2:=(y).Q}{x:=<a>.P{c2:=(y).Q}})")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6569
    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
  6570
    apply(simp add: 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
  6571
    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
  6572
    apply(simp add: fresh_fun_simp_OrR2 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
  6573
    apply(simp_all add: fresh_atm subst_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
  6574
    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
  6575
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6576
    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
  6577
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6578
  case (ImpR z c M 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
  6579
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6580
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6581
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{d:=(y).Q},a,P{d:=(y).Q},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
  6582
                                                     M{d:=(y).Q}{x:=<a>.P{d:=(y).Q}})")
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6583
    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
  6584
    apply(simp add: 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
  6585
    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
  6586
    apply(simp add: fresh_fun_simp_ImpR 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
  6587
    apply(simp_all add: fresh_atm subst_fresh forget 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
  6588
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6589
    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
  6590
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6591
  case (ImpL c M z 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
  6592
  then show ?case  
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6593
    apply(auto simp: fresh_prod fresh_atm subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6594
    apply(subgoal_tac "\<exists>z'::name. z'\<sharp>(P,P{b:=(y).Q},M{u:=<a>.P},N{u:=<a>.P},y,Q,
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6595
                        M{b:=(y).Q}{u:=<a>.P{b:=(y).Q}},N{b:=(y).Q}{u:=<a>.P{b:=(y).Q}},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
  6596
    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
  6597
    apply(simp add: 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
  6598
    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
  6599
    apply(simp add: fresh_fun_simp_ImpL 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
  6600
    apply(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
  6601
    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
  6602
    apply(rule substc.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
  6603
    apply(simp_all add: fresh_atm subst_fresh 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
  6604
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6605
    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
  6606
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
  6607
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6608
lemma subst_subst2:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6609
  assumes a: "a\<sharp>(b,P,N)" "x\<sharp>(y,P,M)" "b\<sharp>(M,N)" "y\<sharp>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
  6610
  shows "M{a:=(x).N}{y:=<b>.P} = M{y:=<b>.P}{a:=(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
  6611
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
  6612
proof(nominal_induct M avoiding: a x N y b P 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
  6613
  case (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
  6614
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6615
    by (auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6616
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6617
  case (Cut d M' u 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
  6618
  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
  6619
    apply(simp add: fresh_atm fresh_prod trm.inject 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
  6620
    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
  6621
    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
  6622
    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
  6623
    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
  6624
    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
  6625
    apply(simp add: abs_fresh subst_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
  6626
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6627
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6628
    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
  6629
    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
  6630
    apply(case_tac "a\<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
  6631
    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
  6632
    apply(simp add: not_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
  6633
    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
  6634
    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
  6635
    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
  6636
    apply(simp add: abs_fresh subst_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
  6637
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6638
    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
  6639
    apply(case_tac "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
  6640
    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
  6641
    apply(simp add: not_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
  6642
    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
  6643
    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
  6644
    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
  6645
    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
  6646
    apply(simp add: subst_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
  6647
    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
  6648
    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
  6649
    apply(case_tac "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
  6650
    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
  6651
    apply(simp add: not_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
  6652
    apply(case_tac "a\<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
  6653
    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
  6654
    apply(simp add: not_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
  6655
    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
  6656
    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
  6657
    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
  6658
    apply(simp add: 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
  6659
    apply(simp add: subst_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
  6660
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6661
    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
  6662
    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
  6663
    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
  6664
    apply(simp add: subst_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
  6665
    apply(simp add: subst_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
  6666
    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
  6667
    apply(case_tac "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
  6668
    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
  6669
    apply(simp add: not_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
  6670
    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
  6671
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6672
  case (NotR z M' 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
  6673
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6674
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6675
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(y,P,N,N{y:=<b>.P},M'{d:=(x).N},M'{y:=<b>.P}{d:=(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
  6676
    apply(erule exE, simp add: 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
  6677
    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
  6678
    apply(simp add: fresh_fun_simp_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
  6679
    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
  6680
    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
  6681
    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
  6682
    apply(simp add: fresh_prod subst_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
  6683
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6684
    apply(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
  6685
    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
  6686
    apply(rule substn.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
  6687
    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
  6688
    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
  6689
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6690
    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
  6691
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6692
  case (NotL d 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
  6693
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6694
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6695
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(z,y,P,N,N{y:=<b>.P},M'{y:=<b>.P},M'{y:=<b>.P}{a:=(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
  6696
    apply(erule exE, simp add: 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
  6697
    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
  6698
    apply(simp add: fresh_fun_simp_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
  6699
    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
  6700
    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
  6701
    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
  6702
    apply(simp add: abs_fresh subst_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
  6703
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6704
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6705
    apply(simp add: trm.inject alpha 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
  6706
    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
  6707
    apply(rule substc.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
  6708
    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
  6709
    apply(simp add: fresh_atm 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
  6710
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6711
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6712
    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
  6713
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6714
  case (AndR d M' e M'' f) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6715
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6716
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6717
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,b,d,e,N,N{y:=<b>.P},M'{f:=(x).N},M''{f:=(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
  6718
                  M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}},M''{y:=<b>.P}{f:=(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
  6719
    apply(erule exE, simp add: 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
  6720
    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
  6721
    apply(simp add: fresh_fun_simp_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
  6722
    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
  6723
    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
  6724
    apply(simp add: abs_fresh subst_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
  6725
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6726
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6727
    apply(simp add: trm.inject alpha 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
  6728
    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
  6729
    apply(rule substn.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
  6730
    apply(simp add: fresh_prod fresh_atm 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
  6731
    apply(simp add: fresh_atm 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
  6732
    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
  6733
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6734
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6735
    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
  6736
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6737
  case (AndL1 z M' 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
  6738
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6739
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6740
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,b,z,u,x,N,M'{y:=<b>.P},M'{y:=<b>.P}{a:=(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
  6741
    apply(erule exE, simp add: 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
  6742
    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
  6743
    apply(simp add: fresh_fun_simp_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
  6744
    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
  6745
    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
  6746
    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
  6747
    apply(simp add: abs_fresh subst_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
  6748
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6749
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6750
    apply(simp add: trm.inject alpha 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
  6751
    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
  6752
    apply(rule substc.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
  6753
    apply(simp add: fresh_prod fresh_atm 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
  6754
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6755
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6756
    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
  6757
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6758
  case (AndL2 z M' 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
  6759
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6760
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6761
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,b,z,u,x,N,M'{y:=<b>.P},M'{y:=<b>.P}{a:=(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
  6762
    apply(erule exE, simp add: 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
  6763
    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
  6764
    apply(simp add: fresh_fun_simp_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
  6765
    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
  6766
    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
  6767
    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
  6768
    apply(simp add: abs_fresh subst_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
  6769
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6770
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6771
    apply(simp add: trm.inject alpha 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
  6772
    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
  6773
    apply(rule substc.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
  6774
    apply(simp add: fresh_prod fresh_atm 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
  6775
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6776
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6777
    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
  6778
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6779
  case (OrL u M' v M'' w) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6780
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6781
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6782
    apply(subgoal_tac "\<exists>z'::name. z'\<sharp>(P,b,u,w,v,N,N{y:=<b>.P},M'{y:=<b>.P},M''{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
  6783
                  M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}},M''{y:=<b>.P}{a:=(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
  6784
    apply(erule exE, simp add: 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
  6785
    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
  6786
    apply(simp add: fresh_fun_simp_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
  6787
    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
  6788
    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
  6789
    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
  6790
    apply(simp add: abs_fresh subst_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
  6791
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6792
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6793
    apply(simp add: trm.inject alpha 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
  6794
    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
  6795
    apply(rule substc.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
  6796
    apply(simp add: fresh_prod fresh_atm 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
  6797
    apply(simp add: fresh_atm 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
  6798
    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
  6799
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6800
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6801
    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
  6802
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6803
  case (OrR1 e M' f) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6804
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6805
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6806
    apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(P,b,e,f,x,N,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
  6807
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(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
  6808
    apply(erule exE, simp add: 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
  6809
    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
  6810
    apply(simp add: fresh_fun_simp_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
  6811
    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
  6812
    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
  6813
    apply(simp add: abs_fresh subst_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
  6814
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6815
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6816
    apply(simp add: trm.inject alpha 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
  6817
    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
  6818
    apply(rule substn.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
  6819
    apply(simp add: fresh_prod fresh_atm 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
  6820
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6821
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6822
    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
  6823
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6824
  case (OrR2 e M' f) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6825
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6826
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6827
    apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(P,b,e,f,x,N,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
  6828
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(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
  6829
    apply(erule exE, simp add: 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
  6830
    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
  6831
    apply(simp add: fresh_fun_simp_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
  6832
    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
  6833
    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
  6834
    apply(simp add: abs_fresh subst_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
  6835
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6836
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6837
    apply(simp add: trm.inject alpha 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
  6838
    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
  6839
    apply(rule substn.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
  6840
    apply(simp add: fresh_prod fresh_atm 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
  6841
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6842
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6843
    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
  6844
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6845
  case (ImpR x e M' f) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6846
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6847
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6848
    apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(P,b,e,f,x,N,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
  6849
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(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
  6850
    apply(erule exE, simp add: 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
  6851
    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
  6852
    apply(simp add: fresh_fun_simp_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
  6853
    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
  6854
    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
  6855
    apply(simp add: abs_fresh subst_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
  6856
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6857
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6858
    apply(simp add: trm.inject alpha 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
  6859
    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
  6860
    apply(rule substn.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
  6861
    apply(simp add: fresh_prod fresh_atm 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
  6862
    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
  6863
    apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_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
  6864
    apply(rule exists_fresh'(2)[OF fs_coname1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6865
    apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_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
  6866
    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
  6867
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6868
  case (ImpL e M' v M'' w) 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6869
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6870
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6871
    apply(subgoal_tac "\<exists>z'::name. z'\<sharp>(P,b,e,w,v,N,N{y:=<b>.P},M'{w:=<b>.P},M''{w:=<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
  6872
                  M'{w:=<b>.P}{a:=(x).N{w:=<b>.P}},M''{w:=<b>.P}{a:=(x).N{w:=<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
  6873
    apply(erule exE, simp add: 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
  6874
    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
  6875
    apply(simp add: fresh_fun_simp_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
  6876
    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
  6877
    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
  6878
    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
  6879
    apply(simp add: abs_fresh subst_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
  6880
    apply(simp add: fresh_prod subst_fresh fresh_atm 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
  6881
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6882
    apply(simp add: trm.inject alpha 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
  6883
    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
  6884
    apply(rule substc.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
  6885
    apply(simp add: fresh_prod fresh_atm 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
  6886
    apply(simp add: fresh_atm 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
  6887
    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
  6888
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6889
    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
  6890
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
  6891
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6892
lemma subst_subst3:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6893
  assumes a: "a\<sharp>(P,N,c)" "c\<sharp>(M,N)" "x\<sharp>(y,P,M)" "y\<sharp>(P,x)" "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
  6894
  shows "N{x:=<a>.M}{y:=<c>.P} = N{y:=<c>.P}{x:=<a>.(M{y:=<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
  6895
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
  6896
proof(nominal_induct N avoiding: x y a c M P 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
  6897
  case (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
  6898
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6899
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6900
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6901
  case (Cut d M' u 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
  6902
  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
  6903
    apply(simp add: fresh_atm fresh_prod trm.inject 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
  6904
    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
  6905
    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
  6906
    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
  6907
    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
  6908
    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
  6909
    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
  6910
    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
  6911
    apply(simp add: fresh_prod subst_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
  6912
    apply(subgoal_tac "P \<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
  6913
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6914
    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
  6915
    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
  6916
    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
  6917
    apply(case_tac "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
  6918
    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
  6919
    apply(simp add: not_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
  6920
    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
  6921
    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
  6922
    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
  6923
    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
  6924
    apply(simp add: fresh_prod subst_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
  6925
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6926
    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
  6927
    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
  6928
    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
  6929
    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
  6930
    apply(simp add: fresh_prod subst_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
  6931
    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
  6932
    apply(case_tac "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
  6933
    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
  6934
    apply(simp add: not_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
  6935
    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
  6936
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6937
  case 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
  6938
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6939
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6940
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6941
  case (NotL d M' 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
  6942
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6943
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6944
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  6945
    apply(erule exE, simp add: 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
  6946
    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
  6947
    apply(simp add: fresh_fun_simp_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
  6948
    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
  6949
    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
  6950
    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
  6951
    apply(simp add: fresh_prod subst_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
  6952
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6953
    apply(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
  6954
    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
  6955
    apply(rule substn.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
  6956
    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
  6957
    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
  6958
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6959
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  6960
    apply(erule exE, simp add: 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
  6961
    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
  6962
    apply(simp add: fresh_fun_simp_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
  6963
    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
  6964
    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
  6965
    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
  6966
    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
  6967
    apply(simp add: fresh_atm subst_fresh 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
  6968
    apply(subgoal_tac "P \<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
  6969
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6970
    apply(simp add: forget 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
  6971
    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
  6972
    apply(rule substn.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
  6973
    apply(simp add: fresh_atm 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
  6974
    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
  6975
    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
  6976
    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
  6977
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6978
    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
  6979
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6980
  case 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
  6981
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6982
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6983
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6984
  case (AndL1 u M' 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
  6985
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  6986
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6987
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(u,y,v,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  6988
    apply(erule exE, simp add: 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
  6989
    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
  6990
    apply(simp add: fresh_fun_simp_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
  6991
    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
  6992
    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
  6993
    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
  6994
    apply(simp add: fresh_prod subst_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
  6995
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  6996
    apply(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
  6997
    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
  6998
    apply(rule substn.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
  6999
    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
  7000
    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
  7001
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7002
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,u,v,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  7003
    apply(erule exE, simp add: 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
  7004
    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
  7005
    apply(simp add: fresh_fun_simp_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
  7006
    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
  7007
    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
  7008
    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
  7009
    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
  7010
    apply(simp add: fresh_atm subst_fresh 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
  7011
    apply(subgoal_tac "P \<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
  7012
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7013
    apply(simp add: forget 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
  7014
    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
  7015
    apply(rule substn.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
  7016
    apply(simp add: fresh_atm 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
  7017
    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
  7018
    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
  7019
    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
  7020
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7021
    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
  7022
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7023
  case (AndL2 u M' 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
  7024
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7025
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7026
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(u,y,v,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  7027
    apply(erule exE, simp add: 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
  7028
    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
  7029
    apply(simp add: fresh_fun_simp_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
  7030
    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
  7031
    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
  7032
    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
  7033
    apply(simp add: fresh_prod subst_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
  7034
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7035
    apply(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
  7036
    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
  7037
    apply(rule substn.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
  7038
    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
  7039
    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
  7040
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7041
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,u,v,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  7042
    apply(erule exE, simp add: 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
  7043
    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
  7044
    apply(simp add: fresh_fun_simp_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
  7045
    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
  7046
    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
  7047
    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
  7048
    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
  7049
    apply(simp add: fresh_atm subst_fresh 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
  7050
    apply(subgoal_tac "P \<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
  7051
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7052
    apply(simp add: forget 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
  7053
    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
  7054
    apply(rule substn.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
  7055
    apply(simp add: fresh_atm 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
  7056
    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
  7057
    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
  7058
    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
  7059
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7060
    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
  7061
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7062
  case 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
  7063
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7064
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7065
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7066
  case 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
  7067
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7068
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7069
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7070
  case (OrL x1 M' x2 M'' 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
  7071
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7072
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7073
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  7074
                                      x1,x2,x3,M''{x:=<a>.M},M''{y:=<c>.P}{x:=<a>.M{y:=<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
  7075
    apply(erule exE, simp add: 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
  7076
    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
  7077
    apply(simp add: fresh_fun_simp_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
  7078
    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
  7079
    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
  7080
    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
  7081
    apply(simp add: fresh_prod subst_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
  7082
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7083
    apply(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
  7084
    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
  7085
    apply(rule substn.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
  7086
    apply(simp add: fresh_prod fresh_atm 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
  7087
    apply(simp add: fresh_prod fresh_atm 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
  7088
    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
  7089
    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
  7090
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7091
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<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
  7092
                                      x1,x2,x3,M''{y:=<c>.P},M''{y:=<c>.P}{x:=<a>.M{y:=<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
  7093
    apply(erule exE, simp add: 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
  7094
    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
  7095
    apply(simp add: fresh_fun_simp_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
  7096
    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
  7097
    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
  7098
    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
  7099
    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
  7100
    apply(simp add: fresh_atm subst_fresh 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
  7101
    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
  7102
    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
  7103
    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
  7104
    apply(simp add: forget 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
  7105
    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
  7106
    apply(rule substn.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
  7107
    apply(simp add: fresh_atm 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
  7108
    apply(simp add: fresh_atm 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
  7109
    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
  7110
    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
  7111
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7112
    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
  7113
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7114
  case 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
  7115
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7116
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7117
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7118
  case (ImpL d M' x1 M'' x2)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7119
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7120
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7121
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,P,M,M{y:=<c>.P},M'{x2:=<a>.M},M'{y:=<c>.P}{x2:=<a>.M{y:=<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
  7122
                                      x1,x2,M''{x2:=<a>.M},M''{y:=<c>.P}{x2:=<a>.M{y:=<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
  7123
    apply(erule exE, simp add: 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
  7124
    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
  7125
    apply(simp add: fresh_fun_simp_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
  7126
    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
  7127
    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
  7128
    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
  7129
    apply(simp add: fresh_prod subst_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
  7130
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7131
    apply(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
  7132
    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
  7133
    apply(rule substn.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
  7134
    apply(simp add: fresh_prod fresh_atm 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
  7135
    apply(simp add: fresh_prod fresh_atm 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
  7136
    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
  7137
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7138
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,P,M,M'{x2:=<c>.P},M'{x2:=<c>.P}{x:=<a>.M{x2:=<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
  7139
                                      x1,x2,M''{x2:=<c>.P},M''{x2:=<c>.P}{x:=<a>.M{x2:=<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
  7140
    apply(erule exE, simp add: 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
  7141
    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
  7142
    apply(simp add: fresh_fun_simp_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
  7143
    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
  7144
    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
  7145
    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
  7146
    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
  7147
    apply(simp add: fresh_atm subst_fresh 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
  7148
    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
  7149
    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
  7150
    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
  7151
    apply(simp add: forget 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
  7152
    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
  7153
    apply(rule substn.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
  7154
    apply(simp add: fresh_atm 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
  7155
    apply(simp add: fresh_atm 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
  7156
    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
  7157
    apply(rule exists_fresh'(1)[OF fs_name1])
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7158
    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
  7159
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
  7160
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7161
lemma subst_subst4:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7162
  assumes a: "x\<sharp>(P,N,y)" "y\<sharp>(M,N)" "a\<sharp>(c,P,M)" "c\<sharp>(P,a)" "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
  7163
  shows "N{a:=(x).M}{c:=(y).P} = N{c:=(y).P}{a:=(x).(M{c:=(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
  7164
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
  7165
proof(nominal_induct N avoiding: x y a c M P 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
  7166
  case (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
  7167
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7168
    by (auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7169
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7170
  case (Cut d M' u 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
  7171
  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
  7172
    apply(simp add: fresh_atm fresh_prod trm.inject 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
  7173
    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
  7174
    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
  7175
    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
  7176
    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
  7177
    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
  7178
    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
  7179
    apply(simp add: abs_fresh subst_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
  7180
    apply(simp add: fresh_prod subst_fresh 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
  7181
    apply(subgoal_tac "P \<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
  7182
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7183
    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
  7184
    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
  7185
    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
  7186
    apply(case_tac "a\<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
  7187
    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
  7188
    apply(simp add: not_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
  7189
    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
  7190
    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
  7191
    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
  7192
    apply(simp add: fresh_prod subst_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
  7193
    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
  7194
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7195
    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
  7196
    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
  7197
    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
  7198
    apply(simp add: fresh_prod subst_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
  7199
    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
  7200
    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
  7201
    apply(case_tac "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
  7202
    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
  7203
    apply(simp add: not_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
  7204
    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
  7205
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7206
  case 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
  7207
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7208
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7209
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7210
  case (NotR u M' 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
  7211
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7212
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7213
    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
  7214
    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
  7215
    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
  7216
    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
  7217
    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
  7218
    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
  7219
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7220
    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
  7221
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7222
    apply(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
  7223
    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
  7224
    apply(rule substc.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
  7225
    apply(simp add: fresh_prod fresh_atm)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7226
    apply(auto simp: fresh_atm fresh_prod)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7227
    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
  7228
    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
  7229
    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
  7230
    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
  7231
    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
  7232
    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
  7233
    apply(simp add: fresh_prod fresh_atm 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
  7234
    apply(simp add: abs_fresh subst_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7235
    apply(auto simp: fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7236
    apply(simp add: trm.inject alpha 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
  7237
    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
  7238
    apply(rule substc.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
  7239
    apply(simp add: fresh_atm subst_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7240
    apply(auto simp: fresh_prod fresh_atm) 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7241
    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
  7242
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7243
  case 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
  7244
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7245
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7246
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7247
  case 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
  7248
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7249
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7250
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7251
  case (AndR d M e M' f)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7252
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7253
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7254
    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
  7255
    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
  7256
    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
  7257
    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
  7258
    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
  7259
    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
  7260
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7261
    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
  7262
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7263
    apply(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
  7264
    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
  7265
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7266
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7267
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7268
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7269
    apply(auto simp: fresh_atm fresh_prod)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7270
    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
  7271
    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
  7272
    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
  7273
    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
  7274
    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
  7275
    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
  7276
    apply(simp add: subst_fresh 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
  7277
    apply(simp add: abs_fresh subst_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7278
    apply(auto simp: fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7279
    apply(simp add: trm.inject alpha 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
  7280
    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
  7281
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7282
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7283
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7284
    apply(simp)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7285
    apply(auto simp: fresh_atm fresh_prod)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7286
    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
  7287
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7288
  case 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
  7289
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7290
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7291
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7292
  case (OrR1 d 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
  7293
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7294
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7295
    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
  7296
    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
  7297
    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
  7298
    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
  7299
    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
  7300
    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
  7301
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7302
    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
  7303
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7304
    apply(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
  7305
    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
  7306
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7307
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7308
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7309
    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
  7310
    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
  7311
    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
  7312
    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
  7313
    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
  7314
    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
  7315
    apply(simp add: subst_fresh 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
  7316
    apply(simp add: abs_fresh subst_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7317
    apply(auto simp: fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7318
    apply(simp add: trm.inject alpha 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
  7319
    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
  7320
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7321
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7322
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7323
    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
  7324
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7325
  case (OrR2 d 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
  7326
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7327
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7328
    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
  7329
    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
  7330
    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
  7331
    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
  7332
    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
  7333
    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
  7334
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7335
    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
  7336
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7337
    apply(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
  7338
    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
  7339
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7340
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7341
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7342
    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
  7343
    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
  7344
    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
  7345
    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
  7346
    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
  7347
    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
  7348
    apply(simp add: subst_fresh 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
  7349
    apply(simp add: abs_fresh subst_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7350
    apply(auto simp: fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7351
    apply(simp add: trm.inject alpha 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
  7352
    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
  7353
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7354
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7355
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7356
    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
  7357
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7358
  case 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
  7359
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7360
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7361
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7362
  case (ImpR u d 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
  7363
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7364
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7365
    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
  7366
    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
  7367
    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
  7368
    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
  7369
    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
  7370
    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
  7371
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7372
    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
  7373
    apply(simp)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7374
    apply(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
  7375
    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
  7376
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7377
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7378
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7379
    apply(auto simp: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7380
    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
  7381
    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
  7382
    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
  7383
    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
  7384
    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
  7385
    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
  7386
    apply(simp add: subst_fresh 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
  7387
    apply(simp add: abs_fresh subst_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7388
    apply(auto simp: fresh_atm)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7389
    apply(simp add: trm.inject alpha 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
  7390
    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
  7391
    apply(rule substc.simps)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7392
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7393
    apply(auto simp: fresh_prod fresh_atm subst_fresh)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7394
    apply(auto simp: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1]
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  7395
    apply(auto simp: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7396
    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
  7397
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
  7398
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  7399
end