| author | wenzelm | 
| Thu, 30 May 2013 22:30:38 +0200 | |
| changeset 52264 | cdba0c3cb4c2 | 
| parent 48567 | 3c4d7ff75f01 | 
| child 53015 | a1119cf551e8 | 
| permissions | -rw-r--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 | 1 | theory Class1 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2 | imports "../Nominal" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3 | begin | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5 | section {* Term-Calculus from Urban's PhD *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 9 | text {* types *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 11 | 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 | 12 | 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 | 13 | | 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 | 14 |   | 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 | 15 |   | 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 | 16 |   | 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 | 17 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 18 | 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 | 19 | 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 | 20 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 21 | 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 | 22 | 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 | 23 | "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 | 24 | | "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 | 25 | | "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 | 26 | | "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 | 27 | | "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 | 28 | 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 | 29 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 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 | 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 | 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 | 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 | 35 | 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 | 36 | 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 | 37 | 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 | 38 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 40 | 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 | 41 | 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 | 42 | 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 | 43 | 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 | 44 | by (nominal_induct T rule: ty.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 45 | (auto simp add: fresh_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 | 46 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 47 | text {* terms *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 48 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 49 | 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 | 50 | 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 | 51 |   | 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 | 52 |   | 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 | 53 |   | 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 | 54 |   | 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 | 55 |   | 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 | 56 |   | 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 | 57 |   | 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 | 58 |   | 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 | 59 |   | 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 | 60 |   | 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 | 61 |   | 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 | 62 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* named terms *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 65 | 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 | 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 | text {* conamed terms *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 69 | 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 | 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 | text {* renaming functions *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 73 | 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 | 74 |   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 | 75 | 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 | 76 | "(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 | 77 | | "\<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 | 78 | | "(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 | 79 | | "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 | 80 | | "\<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 | 81 | (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 | 82 | | "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 | 83 | | "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 | 84 | | "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 | 85 | | "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 | 86 | | "\<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 | 87 | | "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 | 88 | (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 | 89 | | "\<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" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(finite_guess)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 91 | apply(rule TrueI)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 92 | apply(simp add: abs_fresh abs_supp fin_supp)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 93 | 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 | 94 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 97 |   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 | 98 | 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 | 99 | "(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 | 100 | | "\<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 | 101 | | "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 | 102 | | "(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 | 103 | | "\<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 | 104 | | "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 | 105 | | "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 | 106 | | "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 | 107 | | "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 | 108 | | "\<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 | 109 | (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 | 110 | | "\<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 | 111 | | "\<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 | 112 | (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)" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 113 | apply(finite_guess)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 114 | apply(rule TrueI)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 115 | apply(simp add: abs_fresh abs_supp fs_name1 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 | 116 | 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 | 117 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 118 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 119 | 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 | 120 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 122 | 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 | 123 | shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>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 | 124 | apply(nominal_induct M avoiding: d e 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 | 125 | apply(auto simp add: fresh_bij eq_bij) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 126 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 128 | 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 | 129 | 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 | 130 | shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>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 | 131 | apply(nominal_induct M avoiding: d e 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 | 132 | apply(auto simp add: fresh_bij eq_bij) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 133 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 136 | 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 | 137 | shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>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 | 138 | apply(nominal_induct M avoiding: x y rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 139 | apply(auto simp add: fresh_bij eq_bij) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 140 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 142 | 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 | 143 | 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 | 144 | shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>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 | 145 | apply(nominal_induct M avoiding: x y rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 146 | apply(auto simp add: fresh_bij eq_bij) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 147 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 149 | 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 | 150 | 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 | 151 | 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 | 152 | 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 | 153 | 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 | 154 | 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 | 155 | by (nominal_induct M avoiding: x y rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 156 | (auto simp add: trm.inject fresh_atm 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 | 157 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 159 | 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 | 160 | 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 | 161 | 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 | 162 | by (nominal_induct M avoiding: a b rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 163 | (auto simp add: trm.inject 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 | 164 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 166 | 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 | 167 | 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 | 168 | by (nominal_induct M avoiding: x y rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 169 | (auto simp add: fresh_atm 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 | 170 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 172 | 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 | 173 | 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 | 174 | by (nominal_induct M avoiding: a b rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 175 | (auto simp add: fresh_atm 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 | 176 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 178 | 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 | 179 | 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 | 180 | by (nominal_induct M avoiding: a b rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 181 | (auto simp add: fresh_atm 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 | 182 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 184 | 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 | 185 | 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 | 186 | by (nominal_induct M avoiding: x y rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 187 | (auto simp add: fresh_atm 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 | 188 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 190 | 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 | 191 | 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 | 192 | by (nominal_induct M avoiding: x z 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 | 193 | (auto simp add: fresh_prod fresh_atm 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 | 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_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 | 196 | 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 | 197 | 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 | 198 | by (nominal_induct M avoiding: a b 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 | 199 | (auto simp add: fresh_prod fresh_atm 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 | 200 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 202 | 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 | 203 | 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 | 204 | 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 | 205 | apply(nominal_induct M avoiding: x x' y rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 206 | apply(auto simp add: abs_fresh fresh_bij fresh_atm fresh_prod fresh_right calc_atm abs_supp fin_supp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 207 | apply(auto 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 | 208 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 209 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 210 | lemma crename_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 | 211 | 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 | 212 | 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 | 213 | using a | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 214 | apply(nominal_induct M avoiding: a a' b rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 215 | apply(auto simp add: abs_fresh fresh_bij fresh_atm fresh_prod fresh_right calc_atm abs_supp fin_supp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 216 | apply(auto 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 | 217 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 219 | 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 | 220 | 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 | 221 | 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 | 222 | 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 | 223 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 224 | lemma 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 | 225 | 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 | 226 | 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 | 227 | 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 | 228 | 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 | 229 | 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 | 230 | have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (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 | 231 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 232 | 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 | 233 | 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 | 234 | 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 | 235 | 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 | 236 | apply(rule nrename.simps) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 237 | 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 | 238 | 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 | 239 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])" 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 | 241 | 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 | 242 | apply(simp add: trm.inject alpha fresh_atm fresh_prod 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 | 243 | 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 | 244 | 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 | 245 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 247 | 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 | 248 | 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 | 249 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 251 | 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 | 252 | 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 | 253 | 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 | 254 | 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 | 255 | 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 | 256 | have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (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 | 257 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 258 | 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 | 259 | 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 | 260 | 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 | 261 | 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 | 262 | 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 | 263 | 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 | 264 | 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 | 265 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> = Cut <a>.(M[b\<turnstile>c>c]) (x).(N[b\<turnstile>c>c])" 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 | 267 | 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 | 268 | apply(simp add: trm.inject alpha fresh_atm fresh_prod 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 | 269 | 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 | 270 | 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 | 271 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 273 | 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 | 274 | 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 | 275 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 277 | 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 | 278 | 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 | 279 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 280 | 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 | 281 | 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 | 282 | 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 | 283 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 285 | 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 | 286 | by (nominal_induct M avoiding: x y rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 287 | (simp_all add: calc_atm fresh_atm trm.inject alpha 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 | 288 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 290 | 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 | 291 | by (nominal_induct M avoiding: a b rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 292 | (simp_all add: calc_atm fresh_atm trm.inject alpha 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 | 293 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 295 | 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 | 296 | 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 | 297 | 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 | 298 | 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 | 299 | 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 | 300 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 301 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 302 | 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 | 303 | 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 | 304 | 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 | 305 | 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 | 306 | 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 | 307 | 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 | 308 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 309 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 310 | text {* substitution functions *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 311 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 312 | 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 | 313 | 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 | 314 | 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 | 315 | 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 | 316 | 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 | 317 | shows "c\<sharp>(pi\<bullet>M)" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 318 | 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 | 319 | 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 | 320 | apply(simp add: 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 | 321 | apply(simp add: at_prm_fresh[OF at_coname_inst] fresh_list_rev) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 325 | 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 | 326 | 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 | 327 | 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 | 328 | 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 | 329 | shows "x\<sharp>(pi\<bullet>M)" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 330 | 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 | 331 | 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 | 332 | apply(simp add: 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 | 333 | apply(simp add: at_prm_fresh[OF at_name_inst] fresh_list_rev) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 334 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 335 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 336 | 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 | 337 | 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 | 338 | shows "fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M 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 | 339 | 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 | 340 | 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 | 341 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 342 | apply(rule pt_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 | 343 | apply(rule 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 | 344 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 345 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,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 | 346 | 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 | 347 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 348 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 349 | 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 | 350 | 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 | 351 | apply(simp add: 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 | 352 | 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 | 353 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 354 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 355 | 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 | 356 | 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 | 357 | 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 | 358 | 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 | 359 | 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 | 360 | 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 | 361 | fresh_fun (pi2\<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 | 362 | 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 | 363 | 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 | 364 | 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 | 365 | apply(auto 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 | 366 | 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 | 367 | 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 | 368 | 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 | 369 | apply(rule 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 | 370 | apply(rule 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 | 371 | apply(assumption) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 372 | apply(assumption) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 373 | apply(rule 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 | 374 | apply(assumption) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 375 | apply(assumption) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 376 | apply(simp add: at_prm_fresh[OF at_name_inst] swap_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 | 377 | 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 | 378 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 379 | 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 | 380 | 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 | 381 | apply(simp add: fresh_fun_simp_NotL 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 | 382 | 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 | 383 | apply(simp add: 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 | 384 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 385 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 386 | 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 | 387 | 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 | 388 | shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M 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 | 389 | 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 | 390 | 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 | 391 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 392 | apply(rule pt_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 | 393 | apply(rule 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 | 394 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 395 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,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 | 396 | 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 | 397 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 398 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 399 | 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 | 400 | 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 | 401 | apply(simp add: 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 | 402 | 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 | 403 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 404 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 406 | 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 | 407 | 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 | 408 | 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 | 409 | 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 | 410 | 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 | 411 | 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 | 412 | 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 | 413 | 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 | 414 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 415 | apply(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 | 416 | 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 | 417 | apply(simp add: fresh_fun_simp_AndL1 at_prm_fresh[OF at_name_inst] swap_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 | 418 | 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 | 419 | apply(simp add: 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 | 420 | 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 | 421 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)") | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 423 | 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 | 424 | apply(simp add: fresh_fun_simp_AndL1 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 | 425 | 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 | 426 | apply(simp add: 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 | 427 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 428 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 429 | 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 | 430 | 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 | 431 | shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M 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 | 432 | 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 | 433 | 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 | 434 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 435 | apply(rule pt_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 | 436 | apply(rule 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 | 437 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 438 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,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 | 439 | 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 | 440 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 441 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 442 | 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 | 443 | 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 | 444 | apply(simp add: 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 | 445 | 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 | 446 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 447 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 448 | 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 | 449 | 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 | 450 | 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 | 451 | 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 | 452 | 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 | 453 | 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 | 454 | fresh_fun (pi2\<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 | 455 | 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 | 456 | apply(perm_simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 457 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 458 | apply(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 | 459 | 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 | 460 | apply(simp add: fresh_fun_simp_AndL2 at_prm_fresh[OF at_name_inst] swap_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 | 461 | 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 | 462 | apply(simp add: 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 | 463 | 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 | 464 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 465 | apply(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 | 466 | 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 | 467 | apply(simp add: fresh_fun_simp_AndL2 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 | 468 | 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 | 469 | apply(simp add: 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 | 470 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 473 | assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>N" "z'\<sharp>u" "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 | 474 | 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'" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 475 | 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 | 476 | 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 | 477 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 478 | apply(rule pt_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 | 479 | apply(rule 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 | 480 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 481 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,x,M,u,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 | 482 | 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 | 483 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 484 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 485 | 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 | 486 | 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 | 487 | apply(simp add: 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 | 488 | 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 | 489 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 490 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 491 | 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 | 492 | 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 | 493 | 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 | 494 | shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).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 | 495 | fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).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 | 496 | and "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).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 | 497 | fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).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 | 498 | 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 | 499 | 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 | 500 | 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)") | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 502 | 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 | 503 | apply(simp add: fresh_fun_simp_OrL at_prm_fresh[OF at_name_inst] swap_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 | 504 | 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 | 505 | apply(simp add: 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 | 506 | 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 | 507 | 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)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 508 | apply(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 | 509 | 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 | 510 | apply(simp add: fresh_fun_simp_OrL 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 | 511 | 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 | 512 | apply(simp add: 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 | 513 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 516 | 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 | 517 | 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'" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 519 | 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 | 520 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 521 | apply(rule pt_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 | 522 | apply(rule 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 | 523 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 524 | apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,x,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 | 525 | 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 | 526 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 527 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 528 | 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 | 529 | 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 | 530 | apply(simp add: 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 | 531 | 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 | 532 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 533 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 534 | 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 | 535 | 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 | 536 | 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 | 537 | 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 | 538 | 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 | 539 | 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 | 540 | fresh_fun (pi2\<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 | 541 | 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 | 542 | 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 | 543 | 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)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 544 | apply(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 | 545 | 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 | 546 | apply(simp add: fresh_fun_simp_ImpL at_prm_fresh[OF at_name_inst] swap_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 | 547 | 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 | 548 | apply(simp add: 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 | 549 | 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 | 550 | 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)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 551 | 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 | 552 | 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 | 553 | apply(simp add: fresh_fun_simp_ImpL 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 | 554 | 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 | 555 | apply(simp add: 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 | 556 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 559 | 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 | 560 | shows "fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P) = 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 | 561 | 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 | 562 | 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 | 563 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 564 | apply(rule pt_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 | 565 | apply(rule 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 | 566 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 567 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,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 | 568 | 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 | 569 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 570 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 571 | 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 | 572 | 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 | 573 | apply(simp add: 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 | 574 | 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 | 575 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 577 | 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 | 578 | 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 | 579 | 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 | 580 | 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 | 581 | 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 | 582 | 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 | 583 | fresh_fun (pi2\<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 | 584 | 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 | 585 | 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 | 586 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi1\<bullet>P,pi1\<bullet>M,pi1)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 587 | apply(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 | 588 | 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 | 589 | apply(simp add: fresh_fun_simp_NotR 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 | 590 | 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 | 591 | apply(simp add: 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 | 592 | 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 | 593 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 594 | 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 | 595 | 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 | 596 | apply(simp add: fresh_fun_simp_NotR at_prm_fresh[OF at_coname_inst] swap_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 | 597 | 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 | 598 | apply(simp add: 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 | 599 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 602 | 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 | 603 | 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" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 605 | 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 | 606 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 607 | apply(rule pt_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 | 608 | apply(rule 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 | 609 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 610 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,b,M,c,N)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 611 | 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 | 612 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 613 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 614 | 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 | 615 | 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 | 616 | apply(simp add: 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 | 617 | 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 | 618 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 621 | 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 | 622 | 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 | 623 | 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 | 624 | 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 | 625 | 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 | 626 | fresh_fun (pi2\<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 | 627 | 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 | 628 | 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 | 629 | 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)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 630 | apply(simp 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 | 631 | 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 | 632 | apply(simp add: fresh_fun_simp_AndR 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 | 633 | 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 | 634 | apply(simp add: 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 | 635 | 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 | 636 | 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)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 637 | apply(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 | 638 | 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 | 639 | apply(simp add: fresh_fun_simp_AndR at_prm_fresh[OF at_coname_inst] swap_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 | 640 | 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 | 641 | apply(simp add: 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 | 642 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 643 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 644 | 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 | 645 | 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 | 646 | shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P) = 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 | 647 | 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 | 648 | 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 | 649 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 650 | apply(rule pt_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 | 651 | apply(rule 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 | 652 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 653 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,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 | 654 | 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 | 655 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 656 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 657 | 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 | 658 | 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 | 659 | apply(simp add: 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 | 660 | 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 | 661 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 662 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 663 | 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 | 664 | 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 | 665 | 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 | 666 | 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 | 667 | 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 | 668 | 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 | 669 | fresh_fun (pi2\<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 | 670 | 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 | 671 | 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 | 672 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 673 | apply(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 | 674 | 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 | 675 | apply(simp add: fresh_fun_simp_OrR1 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 | 676 | 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 | 677 | apply(simp add: 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 | 678 | 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 | 679 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 680 | apply(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 | 681 | 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 | 682 | apply(simp add: fresh_fun_simp_OrR1 at_prm_fresh[OF at_coname_inst] swap_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 | 683 | 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 | 684 | apply(simp add: 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 | 685 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 688 | 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 | 689 | shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P) = 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 | 690 | 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 | 691 | 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 | 692 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 693 | apply(rule pt_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 | 694 | apply(rule 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 | 695 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 696 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,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 | 697 | 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 | 698 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 699 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 700 | 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 | 701 | 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 | 702 | apply(simp add: 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 | 703 | 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 | 704 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 705 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 706 | 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 | 707 | 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 | 708 | 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 | 709 | 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 | 710 | 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 | 711 | 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 | 712 | fresh_fun (pi2\<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 | 713 | 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 | 714 | 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 | 715 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 716 | apply(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 | 717 | 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 | 718 | apply(simp add: fresh_fun_simp_OrR2 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 | 719 | 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 | 720 | apply(simp add: 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 | 721 | 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 | 722 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 723 | apply(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 | 724 | 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 | 725 | apply(simp add: fresh_fun_simp_OrR2 at_prm_fresh[OF at_coname_inst] swap_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 | 726 | 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 | 727 | apply(simp add: 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 | 728 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 729 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 730 | 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 | 731 | 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 | 732 | shows "fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P) = 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 | 733 | 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 | 734 | 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 | 735 | apply(rule fresh_fun_app) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 736 | apply(rule pt_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 | 737 | apply(rule 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 | 738 | apply(finite_guess) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 739 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,y,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 | 740 | 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 | 741 | apply(rule_tac x="n" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 742 | apply(simp add: fresh_prod abs_fresh) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 743 | 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 | 744 | 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 | 745 | apply(simp add: 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 | 746 | 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 | 747 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 748 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 750 | 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 | 751 | 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 | 752 | 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 | 753 | 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 | 754 | 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 | 755 | fresh_fun (pi2\<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 | 756 | 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 | 757 | 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 | 758 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 759 | apply(simp add: fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 760 | apply(auto) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 761 | apply(simp add: fresh_fun_simp_ImpR 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 | 762 | 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 | 763 | apply(simp add: 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 | 764 | 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 | 765 | apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 766 | apply(simp add: fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 767 | apply(auto) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 768 | apply(simp add: fresh_fun_simp_ImpR at_prm_fresh[OF at_coname_inst] swap_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 | 769 | 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 | 770 | apply(simp add: 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 | 771 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 772 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 773 | 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 | 774 |   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 | 775 | 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 | 776 |   "(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 | 777 | | "\<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 | 778 |   (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 | 779 | | "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 | 780 | | "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 | 781 |   (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 | 782 | | "\<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 | 783 |   (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 | 784 | | "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 | 785 |   (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 | 786 |    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 | 787 | | "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 | 788 |   (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 | 789 |    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 | 790 | | "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 | 791 | | "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 | 792 | | "\<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 | 793 |   (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 | 794 |    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 | 795 | | "\<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 | 796 | | "\<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 | 797 |   (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 | 798 |    else 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 | 799 | apply(finite_guess)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 800 | apply(rule TrueI)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 801 | apply(simp add: abs_fresh 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 | 802 | 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 | 803 | apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)", 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 | 804 | 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 | 805 | 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 | 806 | apply(rule exists_fresh', simp add: 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 | 807 | apply(simp add: abs_fresh 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 | 808 | 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 | 809 | apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)", 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 | 810 | 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 | 811 | 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 | 812 | apply(rule exists_fresh', simp add: 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 | 813 | apply(simp add: abs_fresh 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 | 814 | 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 | 815 | apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)", 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 | 816 | 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 | 817 | 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 | 818 | apply(rule exists_fresh', simp add: 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 | 819 | apply(simp add: abs_fresh 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 | 820 | 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 | 821 | apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)", 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 | 822 | 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 | 823 | 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 | 824 | apply(rule exists_fresh', simp add: 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 | 825 | apply(simp add: abs_fresh 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 | 826 | 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 | 827 | apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)", 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 | 828 | 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 | 829 | 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 | 830 | apply(rule exists_fresh', simp add: 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 | 831 | apply(simp add: abs_fresh 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 | 832 | 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 | 833 | apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)", 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 | 834 | 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 | 835 | 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 | 836 | apply(rule exists_fresh', simp add: 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 | 837 | apply(simp add: abs_fresh 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 | 838 | 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 | 839 | apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)", 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 | 840 | 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 | 841 | 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 | 842 | apply(rule exists_fresh', simp add: 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 | 843 | 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 | 844 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 845 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 846 | 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 | 847 |   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 | 848 | 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 | 849 |   "(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 | 850 | | "\<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 | 851 |   (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 | 852 | | "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 | 853 |   (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 | 854 | | "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 | 855 | | "\<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 | 856 |   (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 | 857 |    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 | 858 | | "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 | 859 | | "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 | 860 | | "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 | 861 |   (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 | 862 | | "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 | 863 |   (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 | 864 | | "\<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 | 865 |   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 | 866 | | "\<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 | 867 |   (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 | 868 |    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 | 869 | | "\<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 | 870 |   ImpL <a>.(M{d:=(z).P}) (x).(N{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 | 871 | apply(finite_guess)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 872 | apply(rule TrueI)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 873 | apply(simp add: abs_fresh abs_supp fs_name1 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 | 874 | 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 | 875 | apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)", 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 | 876 | 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 | 877 | 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 | 878 | apply(rule exists_fresh', simp add: 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 | 879 | apply(simp add: abs_fresh 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 | 880 | 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 | 881 | apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)", 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 | 882 | 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 | 883 | 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 | 884 | apply(rule exists_fresh', simp add: 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 | 885 | apply(simp add: abs_fresh 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 | 886 | 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 | 887 | apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)", 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 | 888 | 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 | 889 | 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 | 890 | apply(rule exists_fresh', simp add: 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 | 891 | apply(simp add: abs_fresh 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 | 892 | 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 | 893 | apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)", 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 | 894 | 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 | 895 | 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 | 896 | apply(rule exists_fresh', simp add: 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 | 897 | apply(simp add: abs_fresh 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 | 898 | 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 | 899 | apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)", 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 | 900 | 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 | 901 | 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 | 902 | apply(rule exists_fresh', simp add: 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 | 903 | apply(simp add: abs_fresh 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 | 904 | 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 | 905 | apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)", 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 | 906 | 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 | 907 | apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm 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 | 908 | apply(rule exists_fresh', simp add: 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 | 909 | apply(simp add: abs_fresh 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 | 910 | 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 | 911 | apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)", 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 | 912 | 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 | 913 | 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 | 914 | apply(rule exists_fresh', simp add: 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 | 915 | apply(simp add: abs_fresh 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 | 916 | apply(fresh_guess add: 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 | 917 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 918 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 919 | 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 | 920 | 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 | 921 | 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 | 922 |   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 | 923 |   and   "pi2\<bullet>(M{c:=(x).N}) = (pi2\<bullet>M){(pi2\<bullet>c):=(pi2\<bullet>x).(pi2\<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 | 924 | apply(nominal_induct M avoiding: c x N rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 925 | apply(auto simp add: eq_bij fresh_bij 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 | 926 | 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 | 927 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 928 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 929 | 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 | 930 | 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 | 931 | 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 | 932 |   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 | 933 |   and   "pi2\<bullet>(M{x:=<c>.N}) = (pi2\<bullet>M){(pi2\<bullet>x):=<(pi2\<bullet>c)>.(pi2\<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 | 934 | apply(nominal_induct M avoiding: c x N rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 935 | apply(auto simp add: eq_bij fresh_bij 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 | 936 | 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 | 937 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 938 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 939 | 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 | 940 |   shows "supp (M{y:=<c>.P}) \<subseteq> ((supp M) - {y}) \<union> (supp 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 | 941 | apply(nominal_induct M avoiding: y P 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 | 942 | 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 | 943 | apply(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 944 | 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 | 945 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 946 | 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 | 947 | 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 | 948 | 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 | 949 | 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 | 950 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 951 | apply(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 | 952 | 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 | 953 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 954 | 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 | 955 | 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 | 956 | 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 | 957 | 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 | 958 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 959 | apply(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 | 960 | 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 | 961 | 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 | 962 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 963 | 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 | 964 | 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 | 965 | 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 | 966 | 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 | 967 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 968 | apply(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 | 969 | 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 | 970 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 971 | 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 | 972 | 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 | 973 | 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 | 974 | 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 | 975 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 976 | apply(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 | 977 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 978 | 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 | 979 | 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 | 980 | 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 | 981 | 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 | 982 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 983 | apply(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 | 984 | 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 | 985 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 986 | 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 | 987 | 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 | 988 | 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 | 989 | 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 | 990 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 991 | apply(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 | 992 | 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 | 993 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 994 | 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 | 995 | 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 | 996 | 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 | 997 | 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 | 998 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 999 | apply(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 | 1000 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1001 | 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 | 1002 | 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 | 1003 | 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 | 1004 | 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 | 1005 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1006 | apply(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 | 1007 | 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 | 1008 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1009 | 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 | 1010 | 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 | 1011 | 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 | 1012 | 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 | 1013 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1014 | apply(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 | 1015 | 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 | 1016 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1017 | 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 | 1018 | 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 | 1019 | 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 | 1020 | 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 | 1021 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1022 | apply(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 | 1023 | 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 | 1024 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1025 | 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 | 1026 | 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 | 1027 | 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 | 1028 | 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 | 1029 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1030 | apply(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 | 1031 | 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 | 1032 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1033 | 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 | 1034 | 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 | 1035 | 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 | 1036 | 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 | 1037 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1038 | 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 | 1039 | 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 | 1040 | 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 | 1041 | 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 | 1042 | 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 | 1043 | 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 | 1044 | 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 | 1045 | 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 | 1046 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1047 | apply(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 | 1048 | 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 | 1049 | 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 | 1050 | 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 | 1051 | 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 | 1052 | 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 | 1053 | 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 | 1054 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1055 | apply(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 | 1056 | 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 | 1057 | 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 | 1058 | 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 | 1059 | 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 | 1060 | 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 | 1061 | 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 | 1062 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1063 | 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 | 1064 | 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 | 1065 | 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 | 1066 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1067 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1068 | 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 | 1069 |   shows "supp (M{y:=<c>.P}) \<subseteq> supp (M) \<union> ((supp 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 | 1070 | apply(nominal_induct M avoiding: y P 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 | 1071 | 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 | 1072 | apply(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1073 | 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 | 1074 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1075 | 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 | 1076 | 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 | 1077 | 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 | 1078 | 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 | 1079 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1080 | apply(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 | 1081 | 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 | 1082 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1083 | 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 | 1084 | 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 | 1085 | 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 | 1086 | 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 | 1087 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1088 | 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 | 1089 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1090 | 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 | 1091 | 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 | 1092 | 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 | 1093 | 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 | 1094 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1095 | apply(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 | 1096 | 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 | 1097 | 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 | 1098 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1099 | 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 | 1100 | 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 | 1101 | 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 | 1102 | 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 | 1103 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1104 | apply(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 | 1105 | 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 | 1106 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1107 | 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 | 1108 | 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 | 1109 | 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 | 1110 | 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 | 1111 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1112 | apply(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 | 1113 | 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 | 1114 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1115 | 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 | 1116 | 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 | 1117 | 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 | 1118 | 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 | 1119 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1120 | apply(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 | 1121 | 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 | 1122 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1123 | 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 | 1124 | 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 | 1125 | 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 | 1126 | 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 | 1127 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1128 | apply(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 | 1129 | 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 | 1130 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1131 | 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 | 1132 | 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 | 1133 | 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 | 1134 | 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 | 1135 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1136 | apply(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 | 1137 | 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 | 1138 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1139 | 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 | 1140 | 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 | 1141 | 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 | 1142 | 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 | 1143 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1144 | apply(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 | 1145 | 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 | 1146 | 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 | 1147 | 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 | 1148 | 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 | 1149 | 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 | 1150 | 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 | 1151 | 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 | 1152 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1153 | 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 | 1154 | 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 | 1155 | 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 | 1156 | 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 | 1157 | 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 | 1158 | 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 | 1159 | 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 | 1160 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1161 | apply(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 | 1162 | 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 | 1163 | 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 | 1164 | 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 | 1165 | 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 | 1166 | 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 | 1167 | 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 | 1168 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1169 | apply(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 | 1170 | 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 | 1171 | 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 | 1172 | 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 | 1173 | 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 | 1174 | 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 | 1175 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1176 | apply(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 | 1177 | 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 | 1178 | 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 | 1179 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1182 |   shows "supp (M{c:=(x).P}) \<subseteq> ((supp M) - {c}) \<union> (supp 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 | 1183 | apply(nominal_induct M avoiding: x P 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 | 1184 | 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 | 1185 | apply(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1186 | apply(blast)+ | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1187 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1188 | 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 | 1189 | 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 | 1190 | 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 | 1191 | 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 | 1192 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1193 | 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 | 1194 | 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 | 1195 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1196 | 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 | 1197 | 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 | 1198 | 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 | 1199 | 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 | 1200 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1201 | 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 | 1202 | 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 | 1203 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1204 | 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 | 1205 | 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 | 1206 | 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 | 1207 | 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 | 1208 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1209 | 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 | 1210 | 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 | 1211 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1212 | 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 | 1213 | 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 | 1214 | 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 | 1215 | 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 | 1216 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1217 | 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 | 1218 | 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 | 1219 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1220 | 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 | 1221 | 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 | 1222 | 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 | 1223 | 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 | 1224 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1225 | 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 | 1226 | 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 | 1227 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1228 | 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 | 1229 | 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 | 1230 | 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 | 1231 | 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 | 1232 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1233 | 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 | 1234 | 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 | 1235 | 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 | 1236 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1237 | 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 | 1238 | 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 | 1239 | 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 | 1240 | 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 | 1241 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1242 | 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 | 1243 | 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 | 1244 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1245 | 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 | 1246 | 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 | 1247 | 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 | 1248 | 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 | 1249 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1250 | 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 | 1251 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1252 | 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 | 1253 | 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 | 1254 | 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 | 1255 | 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 | 1256 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1257 | 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 | 1258 | 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 | 1259 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1260 | 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 | 1261 | 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 | 1262 | 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 | 1263 | 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 | 1264 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1265 | 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 | 1266 | 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 | 1267 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1268 | 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 | 1269 | 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 | 1270 | 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 | 1271 | 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 | 1272 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1273 | 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 | 1274 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1275 | 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 | 1276 | 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 | 1277 | 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 | 1278 | 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 | 1279 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1280 | 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 | 1281 | 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 | 1282 | 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 | 1283 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1284 | 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 | 1285 | 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 | 1286 | 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 | 1287 | 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 | 1288 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1289 | 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 | 1290 | 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 | 1291 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1292 | 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 | 1293 | 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 | 1294 | 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 | 1295 | 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 | 1296 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1297 | 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 | 1298 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1299 | 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 | 1300 | 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 | 1301 | 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 | 1302 | 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 | 1303 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1304 | 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 | 1305 | 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 | 1306 | 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 | 1307 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1309 | 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 | 1310 |   shows "supp (M{c:=(x).P}) \<subseteq> (supp M) \<union> ((supp 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 | 1311 | apply(nominal_induct M avoiding: x P 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 | 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(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1314 | 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 | 1315 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1316 | 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 | 1317 | 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 | 1318 | 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 | 1319 | 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 | 1320 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1321 | 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 | 1322 | 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 | 1323 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1324 | 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 | 1325 | 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 | 1326 | 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 | 1327 | 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 | 1328 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1329 | 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 | 1330 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1331 | 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 | 1332 | 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 | 1333 | 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 | 1334 | 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 | 1335 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1336 | 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 | 1337 | 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 | 1338 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1339 | 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 | 1340 | 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 | 1341 | 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 | 1342 | 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 | 1343 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1344 | 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 | 1345 | 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 | 1346 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1347 | 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 | 1348 | 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 | 1349 | 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 | 1350 | 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 | 1351 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1352 | 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 | 1353 | 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 | 1354 | 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 | 1355 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1356 | 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 | 1357 | 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 | 1358 | 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 | 1359 | 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 | 1360 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1361 | 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 | 1362 | 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 | 1363 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1364 | 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 | 1365 | 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 | 1366 | 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 | 1367 | 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 | 1368 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1369 | 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 | 1370 | 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 | 1371 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1372 | 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 | 1373 | 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 | 1374 | 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 | 1375 | 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 | 1376 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1377 | 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 | 1378 | 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 | 1379 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1380 | 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 | 1381 | 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 | 1382 | 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 | 1383 | 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 | 1384 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1385 | 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 | 1386 | 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 | 1387 | 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 | 1388 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1389 | 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 | 1390 | 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 | 1391 | 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 | 1392 | 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 | 1393 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1394 | 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 | 1395 | 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 | 1396 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1397 | 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 | 1398 | 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 | 1399 | 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 | 1400 | 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 | 1401 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1402 | 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 | 1403 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1404 | 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 | 1405 | 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 | 1406 | 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 | 1407 | 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 | 1408 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1409 | 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 | 1410 | 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 | 1411 | 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 | 1412 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1413 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1414 | lemma 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 | 1415 |   shows "(supp M - {y}) \<subseteq> supp (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 | 1416 | apply(nominal_induct M avoiding: y P 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 | 1417 | 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 | 1418 | apply(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1419 | 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 | 1420 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1421 | 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 | 1422 | 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 | 1423 | 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 | 1424 | 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 | 1425 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1427 | 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 | 1428 | 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 | 1429 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1430 | 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 | 1431 | 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 | 1432 | 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 | 1433 | 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 | 1434 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1436 | 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 | 1437 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1438 | 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 | 1439 | 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 | 1440 | 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 | 1441 | 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 | 1442 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1444 | 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 | 1445 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1446 | 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 | 1447 | 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 | 1448 | 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 | 1449 | 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 | 1450 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1452 | 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 | 1453 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1454 | 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 | 1455 | 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 | 1456 | 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 | 1457 | 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 | 1458 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1460 | 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 | 1461 | 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 | 1462 | 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 | 1463 | 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 | 1464 | 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 | 1465 | 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 | 1466 | 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 | 1467 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1468 | 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 | 1469 | 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 | 1470 | 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 | 1471 | 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 | 1472 | 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 | 1473 | 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 | 1474 | 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 | 1475 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1476 | apply(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 | 1477 | 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 | 1478 | 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 | 1479 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1482 |   shows "(supp M) \<subseteq> ((supp (M{y:=<c>.P}))::coname set)"
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1483 | apply(nominal_induct M avoiding: y P 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 | 1484 | 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 | 1485 | apply(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1486 | 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 | 1487 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1488 | 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 | 1489 | 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 | 1490 | 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 | 1491 | 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 | 1492 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1494 | 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 | 1495 | 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 | 1496 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1497 | 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 | 1498 | 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 | 1499 | 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 | 1500 | 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 | 1501 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1503 | 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 | 1504 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<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 | 1505 | 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 | 1506 | 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 | 1507 | 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 | 1508 | 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 | 1509 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1510 | apply(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 | 1511 | 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 | 1512 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1513 | 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 | 1514 | 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 | 1515 | 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 | 1516 | 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 | 1517 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1518 | apply(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 | 1519 | 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 | 1520 | apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<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 | 1521 | 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 | 1522 | 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 | 1523 | 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 | 1524 | 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 | 1525 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1526 | apply(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 | 1527 | 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 | 1528 | 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 | 1529 | 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 | 1530 | 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 | 1531 | 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 | 1532 | 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 | 1533 | 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 | 1534 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1535 | apply(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 | 1536 | 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 | 1537 | 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 | 1538 | 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 | 1539 | 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 | 1540 | 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 | 1541 | 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 | 1542 | apply(simp add: fresh_def abs_supp trm.supp supp_atm) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1544 | 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 | 1545 | 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 | 1546 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1549 |   shows "(supp M - {c}) \<subseteq>  supp (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 | 1550 | apply(nominal_induct M avoiding: x P 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 | 1551 | 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 | 1552 | apply(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1553 | 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 | 1554 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1555 | 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 | 1556 | 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 | 1557 | 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 | 1558 | 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 | 1559 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1560 | 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 | 1561 | 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 | 1562 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1563 | 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 | 1564 | 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 | 1565 | 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 | 1566 | 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 | 1567 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1568 | 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 | 1569 | 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 | 1570 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1571 | 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 | 1572 | 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 | 1573 | 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 | 1574 | 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 | 1575 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1576 | 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 | 1577 | 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 | 1578 | 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 | 1579 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1580 | 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 | 1581 | 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 | 1582 | 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 | 1583 | 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 | 1584 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1585 | 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 | 1586 | 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 | 1587 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1588 | 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 | 1589 | 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 | 1590 | 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 | 1591 | 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 | 1592 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1593 | 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 | 1594 | 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 | 1595 | 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 | 1596 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1597 | 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 | 1598 | 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 | 1599 | 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 | 1600 | 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 | 1601 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1602 | 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 | 1603 | 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 | 1604 | 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 | 1605 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1608 |   shows "(supp M) \<subseteq> ((supp (M{c:=(x).P}))::name set)"
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1609 | apply(nominal_induct M avoiding: x P 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 | 1610 | 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 | 1611 | apply(auto simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1612 | 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 | 1613 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1614 | 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 | 1615 | 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 | 1616 | 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 | 1617 | 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 | 1618 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1619 | 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 | 1620 | 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 | 1621 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1622 | 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 | 1623 | 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 | 1624 | 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 | 1625 | 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 | 1626 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1627 | 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 | 1628 | 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 | 1629 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1630 | 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 | 1631 | 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 | 1632 | 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 | 1633 | 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 | 1634 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1635 | 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 | 1636 | 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 | 1637 | 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 | 1638 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1639 | 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 | 1640 | 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 | 1641 | 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 | 1642 | 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 | 1643 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1644 | 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 | 1645 | 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 | 1646 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1647 | 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 | 1648 | 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 | 1649 | 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 | 1650 | 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 | 1651 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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 | 1652 | 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 | 1653 | 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 | 1654 | 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 | 1655 | apply(subgoal_tac "\<exists>x'::coname. x'\<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 | 1656 | 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 | 1657 | 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 | 1658 | 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 | 1659 | 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 | 1660 | apply(simp add: fresh_def abs_supp trm.supp supp_atm 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(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 | 1662 | 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 | 1663 | 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 | 1664 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1667 | 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 | 1668 | 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 | 1669 | 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 | 1670 | 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 | 1671 |   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 | 1672 |   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 | 1673 |   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 | 1674 |   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 | 1675 |   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 | 1676 |   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 | 1677 |   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 | 1678 |   and   "b\<sharp>(M,P) \<Longrightarrow> b\<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 | 1679 | 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 | 1680 | apply(insert subst_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 | 1681 | apply(simp_all add: fresh_def supp_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1683 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1686 |   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 | 1687 |   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 | 1688 | 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 | 1689 | apply(auto simp add: fresh_atm 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 | 1690 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1693 | 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 | 1694 |   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 | 1695 | 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 | 1696 | 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 | 1697 | 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 | 1698 | then show ?case by (auto simp add: fresh_prod fresh_atm calc_atm 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 | 1699 | 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 | 1700 | 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 | 1701 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_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 | 1702 | 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 | 1703 | case (NotR y 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 | 1704 | 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 | 1705 | apply(auto simp add: calc_atm trm.inject alpha fresh_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 | 1706 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{d:=(x).N},([(c,d)]\<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 | 1707 | 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 | 1708 | 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 | 1709 | 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 | 1710 | 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 | 1711 | 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 | 1712 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1714 | 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 | 1715 | 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 | 1716 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 1718 | apply(simp add: fresh_prod calc_atm 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 | 1719 | apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh 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 | 1720 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{c3:=(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 | 1721 |                   M'{c3:=(x).N},c1,c2,c3,([(c,c3)]\<bullet>M){c:=(x).N},([(c,c3)]\<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 | 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_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 | 1725 | apply (auto simp add: calc_atm trm.inject alpha 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 | 1726 | 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 | 1727 | apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh 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 | 1728 | apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh 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 | 1729 | apply(auto 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 | 1730 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1732 | 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 | 1733 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_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 | 1734 | 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 | 1735 | 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 | 1736 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_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 | 1737 | 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 | 1738 | 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 | 1739 | 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 | 1740 | apply(auto simp add: calc_atm trm.inject alpha fresh_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 | 1741 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},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 | 1742 | 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 | 1743 | 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 | 1744 | 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 | 1745 | 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 | 1746 | 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 | 1747 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1749 | 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 | 1750 | 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 | 1751 | apply(auto simp add: calc_atm trm.inject alpha fresh_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 | 1752 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},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 | 1753 | 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 | 1754 | 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 | 1755 | 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 | 1756 | 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 | 1757 | 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 | 1758 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1760 | 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 | 1761 | 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 | 1762 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1763 | 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 | 1764 | 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 | 1765 | 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 | 1766 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1767 | 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 | 1768 | case (ImpR y 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 | 1769 | 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 | 1770 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1771 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},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 | 1772 | 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 | 1773 | 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 | 1774 | 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 | 1775 | 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 | 1776 | 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 | 1777 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1779 | 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 | 1780 | 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 | 1781 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1782 | 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 | 1783 | 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 | 1784 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1786 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1788 | 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 | 1789 |   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 | 1790 | 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 | 1791 | 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 | 1792 | 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 | 1793 | 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 | 1794 | by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap 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 | 1795 | 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 | 1796 | 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 | 1797 | 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 | 1798 | by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap 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 | 1799 | 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 | 1800 | case (NotR y 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 | 1801 | 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 | 1802 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1803 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{d:=(y).([(y,x)]\<bullet>N)},[(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 | 1804 | 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 | 1805 | 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 | 1806 | 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 | 1807 | 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 | 1808 | 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 | 1809 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1811 | 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 | 1812 | 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 | 1813 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1814 | 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 | 1815 |        "\<exists>a'::coname. a'\<sharp>(N,M{c3:=(y).([(y,x)]\<bullet>N)},M'{c3:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,c1,c2,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 | 1816 | 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 | 1817 | 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 | 1818 | 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 | 1819 | apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh perm_swap 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 | 1820 | 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 | 1821 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1823 | 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 | 1824 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1825 | 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 | 1826 | 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 | 1827 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1828 | 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 | 1829 | 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 | 1830 | 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 | 1831 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1832 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,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 | 1833 | 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 | 1834 | 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 | 1835 | 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 | 1836 | 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 | 1837 | 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 | 1838 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1840 | 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 | 1841 | 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 | 1842 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1843 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,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 | 1844 | 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 | 1845 | 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 | 1846 | 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 | 1847 | 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 | 1848 | 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 | 1849 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1851 | 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 | 1852 | 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 | 1853 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1854 | 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 | 1855 | 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 | 1856 | 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 | 1857 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1858 | 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 | 1859 | case (ImpR y 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 | 1860 | 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 | 1861 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1862 |     apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,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 | 1863 | 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 | 1864 | 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 | 1865 | 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 | 1866 | 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 | 1867 | 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 | 1868 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1870 | 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 | 1871 | 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 | 1872 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left 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 | 1873 | 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 | 1874 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1876 | 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 | 1877 |   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 | 1878 | 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 | 1879 | 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 | 1880 | 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 | 1881 | then show ?case by (auto simp add: fresh_prod fresh_atm calc_atm 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 | 1882 | 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 | 1883 | 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 | 1884 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_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 | 1885 | 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 | 1886 | 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 | 1887 | 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 | 1888 | apply(auto simp add: calc_atm trm.inject alpha fresh_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 | 1889 |     apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,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 | 1890 | 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 | 1891 | 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 | 1892 | 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 | 1893 | 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 | 1894 | 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 | 1895 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1897 | 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 | 1898 | 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 | 1899 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1900 | 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 | 1901 | 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 | 1902 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_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 | 1903 | 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 | 1904 | 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 | 1905 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_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 | 1906 | 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 | 1907 | 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 | 1908 | 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 | 1909 | apply(auto simp add: calc_atm trm.inject alpha fresh_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 | 1910 |     apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.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 | 1911 | 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 | 1912 | 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 | 1913 | 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 | 1914 | 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 | 1915 | 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 | 1916 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1918 | 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 | 1919 | 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 | 1920 | apply(auto simp add: calc_atm trm.inject alpha fresh_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 | 1921 |     apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.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 | 1922 | 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 | 1923 | 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 | 1924 | 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 | 1925 | 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 | 1926 | 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 | 1927 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1929 | 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 | 1930 | 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 | 1931 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1932 | 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 | 1933 |     "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},M'{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N},([(y,x)]\<bullet>M'){y:=<a>.N},x1,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 | 1934 | 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 | 1935 | 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 | 1936 | 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 | 1937 | 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 | 1938 | 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 | 1939 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1941 | 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 | 1942 | 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 | 1943 | by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_left abs_supp fin_supp fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1945 | 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 | 1946 | 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 | 1947 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1948 | 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 | 1949 |     "\<exists>a'::name. a'\<sharp>(N,M{u:=<a>.N},M'{u:=<a>.N},([(y,u)]\<bullet>M){y:=<a>.N},([(y,u)]\<bullet>M'){y:=<a>.N},d,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 | 1950 | 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 | 1951 | 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 | 1952 | 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 | 1953 | 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 | 1954 | 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 | 1955 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1957 | 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 | 1958 | 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 | 1959 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1960 | 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 | 1961 | 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 | 1962 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1964 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1966 | 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 | 1967 |   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 | 1968 | 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 | 1969 | 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 | 1970 | 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 | 1971 | 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 | 1972 | by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap 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 | 1973 | 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 | 1974 | 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 | 1975 | 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 | 1976 | by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap 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 | 1977 | 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 | 1978 | 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 | 1979 | 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 | 1980 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1981 |     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 | 1982 | 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 | 1983 | 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 | 1984 | 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 | 1985 | 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 | 1986 | 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 | 1987 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 1989 | 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 | 1990 | 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 | 1991 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 1992 | 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 | 1993 |        "\<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 | 1994 | 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 | 1995 | 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 | 1996 | 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 | 1997 | apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh perm_swap 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 | 1998 | 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 | 1999 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2001 | 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 | 2002 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2003 | 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 | 2004 | 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 | 2005 | then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2006 | 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 | 2007 | 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 | 2008 | 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 | 2009 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2010 |     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 | 2011 | 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 | 2012 | 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 | 2013 | 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 | 2014 | 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 | 2015 | 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 | 2016 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2018 | 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 | 2019 | 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 | 2020 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2021 |     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 | 2022 | 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 | 2023 | 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 | 2024 | 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 | 2025 | 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 | 2026 | 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 | 2027 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2029 | 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 | 2030 | 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 | 2031 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2032 | 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 | 2033 | 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 | 2034 | 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 | 2035 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2036 | 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 | 2037 | 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 | 2038 | 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 | 2039 | apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2040 |     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 | 2041 | 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 | 2042 | 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 | 2043 | 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 | 2044 | 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 | 2045 | 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 | 2046 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2048 | 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 | 2049 | 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 | 2050 | by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left 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 | 2051 | 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 | 2052 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2054 | 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 | 2055 |   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 | 2056 | 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 | 2057 |   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 | 2058 |   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 | 2059 | 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 | 2060 | 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 | 2061 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2063 | 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 | 2064 |   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 | 2065 | 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 | 2066 |   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 | 2067 |   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 | 2068 | 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 | 2069 | 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 | 2070 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2072 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2074 | 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 | 2075 |   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 | 2076 |   (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 | 2077 | 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 | 2078 | 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 | 2079 | 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 | 2080 | have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (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 | 2081 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 2082 | have eq2: "(M=Ax y a) = (([(a',a)]\<bullet>M)=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 | 2083 | apply(auto 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 | 2084 | 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 | 2085 | 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 | 2086 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2088 | 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 | 2089 |   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 | 2090 |                               else Cut <a'>.(([(a',a)]\<bullet>M){y:=<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 | 2091 | using fs1 fs2 by (auto 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 | 2092 |   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 | 2093 | 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 | 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(simp only: eq2[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 | 2096 | apply(auto 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 | 2097 | 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 | 2098 | 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 | 2099 | 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 | 2100 | 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 | 2101 | 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 | 2102 | 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 | 2103 | 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 | 2104 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2106 | 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 | 2107 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2109 | 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 | 2110 |   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 | 2111 |   (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 | 2112 | 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 | 2113 | 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 | 2114 | 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 | 2115 | have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (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 | 2116 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 2117 | have eq2: "(N=Ax x c) = (([(x',x)]\<bullet>N)=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 | 2118 | apply(auto 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 | 2119 | 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 | 2120 | 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 | 2121 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2123 | 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 | 2124 |   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 | 2125 |                               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 | 2126 | 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 | 2127 |   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 | 2128 | 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 | 2129 | 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 | 2130 | apply(simp only: eq2[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 | 2131 | apply(auto 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 | 2132 | 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 | 2133 | 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 | 2134 | 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 | 2135 | 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 | 2136 | 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 | 2137 | 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 | 2138 | 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 | 2139 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2141 | 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 | 2142 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2144 | 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 | 2145 |   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 | 2146 | 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 | 2147 | 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 | 2148 | 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 | 2149 | 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 | 2150 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2151 | apply(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 | 2152 | 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 | 2153 | 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 | 2154 | apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2155 | 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 | 2156 | 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 | 2157 | 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 | 2158 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2159 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2161 | 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 | 2162 |   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 | 2163 | 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 | 2164 | 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 | 2165 | 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 | 2166 | apply(subgoal_tac "NotR (x).M d = NotR (c).([(c,x)]\<bullet>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 | 2167 | apply(auto simp add: fresh_left calc_atm 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 | 2168 | 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 | 2169 | 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: 
39198diff
changeset | 2170 | 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 | 2171 | 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 | 2172 | 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 | 2173 | 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 | 2174 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2177 | 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 | 2178 |   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 | 2179 | 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 | 2180 | 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 | 2181 | 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 | 2182 | apply(subgoal_tac "NotL <a>.M y = NotL <ca>.([(ca,a)]\<bullet>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 | 2183 | apply(auto simp add: fresh_left calc_atm 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 | 2184 | 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 | 2185 | 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: 
39198diff
changeset | 2186 | 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 | 2187 | 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 | 2188 | 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 | 2189 | 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 | 2190 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2193 | 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 | 2194 |   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 | 2195 | 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 | 2196 | 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 | 2197 | 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 | 2198 | apply(subgoal_tac "AndL1 (x).M y = AndL1 (ca).([(ca,x)]\<bullet>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 | 2199 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2200 | 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 | 2201 | 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: 
39198diff
changeset | 2202 | 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 | 2203 | 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 | 2204 | 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 | 2205 | 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 | 2206 | 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 | 2207 | 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: 
39198diff
changeset | 2208 | 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 | 2209 | 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 | 2210 | 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 | 2211 | 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 | 2212 | 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 | 2213 | 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 | 2214 | 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 | 2215 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2218 | 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 | 2219 |   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 | 2220 | 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 | 2221 | 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 | 2222 | 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 | 2223 | apply(subgoal_tac "AndL2 (x).M y = AndL2 (ca).([(ca,x)]\<bullet>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 | 2224 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2225 | 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 | 2226 | 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: 
39198diff
changeset | 2227 | 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 | 2228 | 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 | 2229 | 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 | 2230 | 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 | 2231 | 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 | 2232 | 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: 
39198diff
changeset | 2233 | 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 | 2234 | 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 | 2235 | 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 | 2236 | 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 | 2237 | 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 | 2238 | 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 | 2239 | 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 | 2240 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2243 | 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 | 2244 |   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 | 2245 | 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 | 2246 | 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 | 2247 | 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 | 2248 | 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 | 2249 | apply(subgoal_tac "AndR <a>.M <b>.N c = AndR <ca>.([(ca,a)]\<bullet>M) <caa>.([(caa,b)]\<bullet>N) c") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2250 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2251 | 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 | 2252 | 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 | 2253 | apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2254 | apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2255 | apply(auto simp add: fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2256 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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: 
39198diff
changeset | 2258 | 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 | 2259 | 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 | 2260 | 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 | 2261 | 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 | 2262 | 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 | 2263 | apply(auto simp add: fresh_left calc_atm abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2264 | 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 | 2265 | apply(auto simp add: fresh_left calc_atm abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2266 | 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 | 2267 | 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 | 2268 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2271 | 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 | 2272 |   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 | 2273 | 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 | 2274 | 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 | 2275 | 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 | 2276 | 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 | 2277 | apply(subgoal_tac "OrL (y).M (z).N x = OrL (ca).([(ca,y)]\<bullet>M) (caa).([(caa,z)]\<bullet>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 | 2278 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2279 | 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 | 2280 | 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 | 2281 | apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2282 | apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2283 | apply(auto simp add: fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2284 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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: 
39198diff
changeset | 2286 | 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 | 2287 | 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 | 2288 | 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 | 2289 | 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 | 2290 | 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 | 2291 | apply(auto simp add: fresh_left calc_atm abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2292 | 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 | 2293 | apply(auto simp add: fresh_left calc_atm abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2294 | 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 | 2295 | 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 | 2296 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2299 | 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 | 2300 |   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 | 2301 | 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 | 2302 | 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 | 2303 | 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 | 2304 | apply(subgoal_tac "OrR1 <a>.M d = OrR1 <c>.([(c,a)]\<bullet>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 | 2305 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2306 | 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: 
39198diff
changeset | 2307 | 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 | 2308 | 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 | 2309 | 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 | 2310 | 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 | 2311 | 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 | 2312 | 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: 
39198diff
changeset | 2313 | 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 | 2314 | 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 | 2315 | 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 | 2316 | 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 | 2317 | 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 | 2318 | 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 | 2319 | 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 | 2320 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2323 | 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 | 2324 |   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 | 2325 | 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 | 2326 | 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 | 2327 | 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 | 2328 | apply(subgoal_tac "OrR2 <a>.M d = OrR2 <c>.([(c,a)]\<bullet>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 | 2329 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2330 | 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: 
39198diff
changeset | 2331 | 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 | 2332 | 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 | 2333 | 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 | 2334 | 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 | 2335 | 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 | 2336 | 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: 
39198diff
changeset | 2337 | 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 | 2338 | 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 | 2339 | 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 | 2340 | 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 | 2341 | 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 | 2342 | 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 | 2343 | 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 | 2344 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2347 | 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 | 2348 |   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 | 2349 | 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 | 2350 | 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 | 2351 | 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 | 2352 | 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 | 2353 | apply(subgoal_tac "ImpR (x).<a>.M d = ImpR (ca).<c>.([(c,a)]\<bullet>[(ca,x)]\<bullet>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 | 2354 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2355 | 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: 
39198diff
changeset | 2356 | 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 | 2357 | 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 | 2358 | 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 | 2359 | 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 | 2360 | 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 | 2361 | 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: 
39198diff
changeset | 2362 | 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 | 2363 | 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 | 2364 | 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 | 2365 | 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 | 2366 | 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 | 2367 | 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 | 2368 | 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 | 2369 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2372 | 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 | 2373 |   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 | 2374 | 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 | 2375 | 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 | 2376 | 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 | 2377 | 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 | 2378 | apply(subgoal_tac "ImpL <a>.M (x).N y = ImpL <ca>.([(ca,a)]\<bullet>M) (caa).([(caa,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 | 2379 | apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2380 | 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: 
39198diff
changeset | 2381 | 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 | 2382 | 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 | 2383 | 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 | 2384 | 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 | 2385 | 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 | 2386 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 2388 | 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 | 2389 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2392 | 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 | 2393 |   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 | 2394 | 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 | 2395 | 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 | 2396 | 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 | 2397 | 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 | 2398 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2401 | 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 | 2402 |   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 | 2403 | 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 | 2404 | 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 | 2405 | 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 | 2406 | 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 | 2407 | 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 | 2408 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2411 | 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 | 2412 |   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 | 2413 | 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 | 2414 | 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 | 2415 | 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 | 2416 | 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 | 2417 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2420 | 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 | 2421 |   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 | 2422 | 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 | 2423 | 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 | 2424 | 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 | 2425 | 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 | 2426 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2429 | 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 | 2430 |   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 | 2431 | 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 | 2432 | apply(nominal_induct M avoiding: x c P a b rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2433 | apply(auto simp add: subst_fresh rename_fresh trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2434 | 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 | 2435 | 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 | 2436 | 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 | 2437 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2438 | apply(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 | 2439 | 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 | 2440 | 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 | 2441 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2442 | apply(simp 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 | 2443 | 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 | 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(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 | 2446 | 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 | 2447 | 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 | 2448 | 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 | 2449 | 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 | 2450 | 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 | 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(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 | 2453 | 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 | 2454 | 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 | 2455 | 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 | 2456 | 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 | 2457 | apply(auto simp add: fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2458 | 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 | 2459 | 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 | 2460 | 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 | 2461 | apply(auto simp add: fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2462 | 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 | 2463 | 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 | 2464 | 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 | 2465 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 2467 | 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 | 2468 | 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 | 2469 | 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 | 2470 | 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 | 2471 | 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 | 2472 | 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 | 2473 | 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 | 2474 | 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 | 2475 | 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 | 2476 | 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 | 2477 | 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 | 2478 | 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 | 2479 | 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 | 2480 | 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 | 2481 | 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 | 2482 | 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 | 2483 | 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 | 2484 | 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 | 2485 | 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 | 2486 | 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 | 2487 | 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 | 2488 | 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 | 2489 | 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 | 2490 | 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 | 2491 | 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 | 2492 | 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 | 2493 | 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 | 2494 | 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 | 2495 | 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 | 2496 | 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 | 2497 |                                   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 | 2498 | 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 | 2499 | 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 | 2500 | 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 | 2501 | 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 | 2502 | 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 | 2503 | 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 | 2504 | 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 | 2505 | 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 | 2506 | 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 | 2507 | 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 | 2508 |                                   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 | 2509 | 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 | 2510 | 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 | 2511 | 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 | 2512 | 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 | 2513 | 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 | 2514 | 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 | 2515 | 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 | 2516 | 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 | 2517 | 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 | 2518 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2521 | 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 | 2522 |   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 | 2523 | 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 | 2524 | apply(nominal_induct M avoiding: x c P a b rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2525 | apply(auto simp add: subst_fresh rename_fresh trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2526 | 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 | 2527 | 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 | 2528 | 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 | 2529 | 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 | 2530 | 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 | 2531 | 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 | 2532 | 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 | 2533 | 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 | 2534 | 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 | 2535 | 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 | 2536 | 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 | 2537 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 2539 | 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 | 2540 | 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 | 2541 | 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 | 2542 | 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 | 2543 | 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 | 2544 | 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 | 2545 | 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 | 2546 | 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 | 2547 | 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 | 2548 | 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 | 2549 |                    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 | 2550 | 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 | 2551 | 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 | 2552 | 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 | 2553 | 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 | 2554 | 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 | 2555 | 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 | 2556 | 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 | 2557 | 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 | 2558 | 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 | 2559 | 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 | 2560 |                          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 | 2561 | 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 | 2562 | 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 | 2563 | 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 | 2564 | 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 | 2565 | 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 | 2566 | 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 | 2567 | 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 | 2568 | 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 | 2569 | 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 | 2570 | 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 | 2571 |                          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 | 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_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 | 2575 | 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 | 2576 | 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 | 2577 | 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 | 2578 | 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 | 2579 | 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 | 2580 | 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 | 2581 | 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 | 2582 |                          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 | 2583 | 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 | 2584 | 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 | 2585 | 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 | 2586 | 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 | 2587 | 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 | 2588 | 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 | 2589 | 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 | 2590 | 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 | 2591 | 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 | 2592 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2595 | 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 | 2596 |   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 | 2597 | 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 | 2598 | apply(nominal_induct M avoiding: x c P y z rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2599 | apply(auto simp add: subst_fresh rename_fresh trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2600 | 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 | 2601 | 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 | 2602 | 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 | 2603 | 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 | 2604 | 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 | 2605 | 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 | 2606 | 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 | 2607 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2608 | apply(drule 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 | 2609 | 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 | 2610 | 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 | 2611 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2612 | apply(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 | 2613 | 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 | 2614 | 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 | 2615 | 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 | 2616 | 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 | 2617 | 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 | 2618 | 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 | 2619 | 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 | 2620 | 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 | 2621 | 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 | 2622 | 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 | 2623 | 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 | 2624 | 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 | 2625 | 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 | 2626 | 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 | 2627 | 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 | 2628 | 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 | 2629 | 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 | 2630 | 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 | 2631 | 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 | 2632 | 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 | 2633 | 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 | 2634 | 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 | 2635 | 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 | 2636 | 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 | 2637 | 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 | 2638 | 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 | 2639 | 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 | 2640 | 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 | 2641 | 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 | 2642 | 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 | 2643 |                                   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 | 2644 | 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 | 2645 | 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 | 2646 | 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 | 2647 | 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 | 2648 | 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 | 2649 | 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 | 2650 | 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 | 2651 | 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 | 2652 | 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 | 2653 | 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 | 2654 |                                   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 | 2655 | 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 | 2656 | 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 | 2657 | 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 | 2658 | 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 | 2659 | 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 | 2660 | 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 | 2661 | 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 | 2662 | 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 | 2663 | 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 | 2664 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2667 | 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 | 2668 |   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 | 2669 | 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 | 2670 | apply(nominal_induct M avoiding: x c P y z rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2671 | apply(auto simp add: subst_fresh rename_fresh trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2672 | 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 | 2673 | 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 | 2674 | 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 | 2675 | 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 | 2676 | 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 | 2677 | 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 | 2678 | 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 | 2679 | 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 | 2680 | 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 | 2681 | 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 | 2682 | 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 | 2683 | 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 | 2684 | 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 | 2685 | 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 | 2686 | 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 | 2687 | 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 | 2688 | 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 | 2689 | 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 | 2690 | 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 | 2691 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2692 | apply(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 | 2693 | 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 | 2694 | 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 | 2695 | 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 | 2696 | 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 | 2697 | 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 | 2698 | 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 | 2699 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 2701 | 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 | 2702 | 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 | 2703 | 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 | 2704 | 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 | 2705 | 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 | 2706 | 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 | 2707 | 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 | 2708 | 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 | 2709 | 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 | 2710 | 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 | 2711 |                    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 | 2712 | 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 | 2713 | 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 | 2714 | 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 | 2715 | 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 | 2716 | 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 | 2717 | 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 | 2718 | 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 | 2719 | 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 | 2720 | 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 | 2721 | 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 | 2722 |                          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 | 2723 | 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 | 2724 | 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 | 2725 | 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 | 2726 | 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 | 2727 | 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 | 2728 | 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 | 2729 | 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 | 2730 | 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 | 2731 | 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 | 2732 | 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 | 2733 |                          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 | 2734 | 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 | 2735 | 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 | 2736 | 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 | 2737 | 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 | 2738 | 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 | 2739 | 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 | 2740 | 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 | 2741 | 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 | 2742 | 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 | 2743 | 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 | 2744 |                          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 | 2745 | 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 | 2746 | 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 | 2747 | 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 | 2748 | 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 | 2749 | 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 | 2750 | 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 | 2751 | 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 | 2752 | 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 | 2753 | 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 | 2754 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2757 | 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 | 2758 |   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 | 2759 | 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 | 2760 | 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 | 2761 | 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 | 2762 | 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 | 2763 | 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 | 2764 |   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 | 2765 | 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 | 2766 | 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 | 2767 | 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 | 2768 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2769 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |    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 | 2771 | 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 | 2772 | 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 | 2773 | 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 | 2774 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2777 | 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 | 2778 | 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 | 2779 | 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 | 2780 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2782 |   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 | 2783 | 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 | 2784 | 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 | 2785 | 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 | 2786 | 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 | 2787 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2789 |   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 | 2790 | 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 | 2791 | 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 | 2792 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2794 | 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 | 2795 |   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 | 2796 | 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 | 2797 | 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 | 2798 | 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 | 2799 | 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 | 2800 | 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 | 2801 | 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 | 2802 |   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 | 2803 | 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 | 2804 | 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 | 2805 | 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 | 2806 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |    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 | 2809 | 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 | 2810 | 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 | 2811 | 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 | 2812 | 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 | 2813 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2815 | 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 | 2816 | 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 | 2817 | 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 | 2818 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2820 |   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 | 2821 | 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 | 2822 | 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 | 2823 | 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 | 2824 | 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 | 2825 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2827 |   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 | 2828 |   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 | 2829 | 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 | 2830 | 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 | 2831 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2833 | 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 | 2834 |   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 | 2835 | 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 | 2836 | 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 | 2837 | 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 | 2838 | 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 | 2839 | 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 | 2840 | 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 | 2841 |   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 | 2842 | 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 | 2843 | 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 | 2844 | 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 | 2845 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2846 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |    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 | 2848 | 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 | 2849 | 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 | 2850 | 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 | 2851 | 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 | 2852 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2854 | 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 | 2855 | 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 | 2856 | 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 | 2857 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2859 |   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 | 2860 | 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 | 2861 | 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 | 2862 | 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 | 2863 | 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 | 2864 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2866 |   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 | 2867 |   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 | 2868 | 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 | 2869 | 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 | 2870 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2872 | 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 | 2873 |   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 | 2874 | 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 | 2875 | 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 | 2876 | 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 | 2877 | 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 | 2878 | 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 | 2879 |   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 | 2880 | 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 | 2881 | 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 | 2882 | 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 | 2883 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |    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 | 2886 | 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 | 2887 | 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 | 2888 | 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 | 2889 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2892 | 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 | 2893 | 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 | 2894 | 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 | 2895 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2897 |   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 | 2898 | 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 | 2899 | 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 | 2900 | 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 | 2901 | 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 | 2902 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 2904 |   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 | 2905 | 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 | 2906 | 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 | 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 | 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 | 2909 | 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 | 2910 | 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 | 2911 | 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 | 2912 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* typing contexts *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 41798 | 2915 | type_synonym ctxtn = "(name\<times>ty) list" | 
| 2916 | 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 | 2917 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2919 | 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 | 2920 | 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 | 2921 | 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 | 2922 | | 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 | 2923 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2925 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2927 | 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 | 2928 | 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 | 2929 | 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 | 2930 | | 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 | 2931 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2933 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2935 | 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 | 2936 | 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 | 2937 | 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 | 2938 | 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 | 2939 | 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 | 2940 | 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 | 2941 | show "a\<sharp>\<Gamma>" by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2942 | 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 | 2943 | show "x\<sharp>\<Delta>" by (induct \<Delta>) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2944 | 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 | 2945 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* cut-reductions *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2947 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2949 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2951 | 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 | 2952 | 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 | 2953 | [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 | 2954 | | [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 | 2955 | | [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 | 2956 | | [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 | 2957 | | [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 | 2958 | | [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 | 2959 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2961 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2963 | 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 | 2964 | 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 | 2965 | 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 | 2966 | apply(erule_tac 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 | 2967 | apply(auto 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 | 2968 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2971 | 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 | 2972 | 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 | 2973 | 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 | 2974 | apply(erule_tac 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 | 2975 | apply(auto 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 | 2976 | 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 | 2977 | 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 | 2978 | 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 | 2979 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2981 | 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 | 2982 | 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 | 2983 | 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 | 2984 | 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 | 2985 | apply(erule_tac 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 | 2986 | apply(auto 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 | 2987 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 2988 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2990 | 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 | 2991 | 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 | 2992 | 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 | 2993 | apply(erule_tac 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 | 2994 | apply(auto 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 | 2995 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 2998 | 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 | 2999 | 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 | 3000 | 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 | 3001 | apply(erule_tac 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 | 3002 | apply(auto 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 | 3003 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3006 | 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 | 3007 | 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 | 3008 | 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 | 3009 | apply(erule_tac 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 | 3010 | apply(auto 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 | 3011 | 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 | 3012 | 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 | 3013 | 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 | 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 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 | 3017 | 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 | 3018 | 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 | 3019 | 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 | 3020 | 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 | 3021 | 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 | 3022 | 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 | 3023 | 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 | 3024 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3026 | 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 | 3027 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3029 | 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 | 3030 | by (induct rule: fin.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 | 3031 | (auto simp add: calc_atm simp add: 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 | 3032 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3034 | 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 | 3035 |   shows "\<not>(fin (M{c:=(y).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 | 3036 | 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 | 3037 | 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 | 3038 | 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 | 3039 | 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 | 3040 | 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 | 3041 | 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 | 3042 | 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 | 3043 | 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 | 3044 | 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 | 3045 | 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 | 3046 | 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 | 3047 | 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 | 3048 | 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 | 3049 | 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 | 3050 | 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 | 3051 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3052 | apply(drule 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 | 3053 | 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 | 3054 | 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 | 3055 | 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 | 3056 | 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 | 3057 | 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 | 3058 | 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 | 3059 | 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 | 3060 | 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 | 3061 | 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 | 3062 | 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 | 3063 | 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 | 3064 | 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 | 3065 | 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 | 3066 | 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 | 3067 | 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 | 3068 | 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 | 3069 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3070 | 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 | 3071 | 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 | 3072 | 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 | 3073 | 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 | 3074 | 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 | 3075 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3076 | 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 | 3077 | 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 | 3078 | 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 | 3079 | 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 | 3080 | 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 | 3081 | 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 | 3082 | 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 | 3083 | 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 | 3084 | 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 | 3085 | 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 | 3086 | 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 | 3087 | 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 | 3088 | 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 | 3089 | 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 | 3090 | 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 | 3091 | 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 | 3092 | 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 | 3093 | 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 | 3094 | 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 | 3095 | 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 | 3096 | 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 | 3097 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3098 | 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 | 3099 | 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 | 3100 | 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 | 3101 | 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 | 3102 | 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 | 3103 | 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 | 3104 | 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 | 3105 | 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 | 3106 | 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 | 3107 | 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 | 3108 | 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 | 3109 | 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 | 3110 | 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 | 3111 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3112 | 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 | 3113 | 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 | 3114 | 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 | 3115 | 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 | 3116 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3119 | 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 | 3120 |   shows "\<not>(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 | 3121 | 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 | 3122 | 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 | 3123 | 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 | 3124 | 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 | 3125 | 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 | 3126 | 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 | 3127 | 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 | 3128 | 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 | 3129 | 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 | 3130 | 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 | 3131 | 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 | 3132 | 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 | 3133 | 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 | 3134 | 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 | 3135 | 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 | 3136 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3138 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3140 | 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 | 3141 | 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 | 3142 | 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 | 3143 | 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 | 3144 | 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 | 3145 | 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 | 3146 | 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 | 3147 | 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 | 3148 | 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 | 3149 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3150 | 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 | 3151 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3152 | apply(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 | 3153 | 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 | 3154 | 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 | 3155 | 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 | 3156 | 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 | 3157 | 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 | 3158 | 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 | 3159 | 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 | 3160 | 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 | 3161 | 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 | 3162 | 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 | 3163 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3164 | 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 | 3165 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3167 | 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 | 3168 | 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 | 3169 | 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 | 3170 | 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 | 3171 | 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 | 3172 | 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 | 3173 | 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 | 3174 | 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 | 3175 | 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 | 3176 | 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 | 3177 | 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 | 3178 | 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 | 3179 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3180 | 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 | 3181 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3183 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3184 | 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 | 3185 | 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 | 3186 | 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 | 3187 | 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 | 3188 | 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 | 3189 | 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 | 3190 | 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 | 3191 | 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 | 3192 | 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 | 3193 | 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 | 3194 | 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 | 3195 | 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 | 3196 | apply(auto simp add: abs_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3197 | 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 | 3198 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3200 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3202 | 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 | 3203 | 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 | 3204 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3206 | 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 | 3207 | 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 | 3208 |   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 | 3209 | 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 | 3210 | 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 | 3211 | 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 | 3212 | 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 | 3213 | 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 | 3214 | 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 | 3215 | 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 | 3216 | 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 | 3217 | 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 | 3218 | 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 | 3219 | 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 | 3220 | 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 | 3221 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3222 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3223 | 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 | 3224 | 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 | 3225 |   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 | 3226 | 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 | 3227 | 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 | 3228 | 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 | 3229 | 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 | 3230 | 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 | 3231 | 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 | 3232 | 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 | 3233 | 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 | 3234 | 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 | 3235 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3236 | apply(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 | 3237 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3238 | apply(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 | 3239 | 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 | 3240 | 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 | 3241 | 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 | 3242 | 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 | 3243 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3245 | 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 | 3246 | 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 | 3247 | 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 | 3248 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3250 | 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 | 3251 | 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 | 3252 | 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 | 3253 | 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 | 3254 | 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 | 3255 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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: 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 | 3257 | 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 | 3258 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3260 | 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 | 3261 | 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 | 3262 | 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 | 3263 | 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 | 3264 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3266 | 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 | 3267 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3271 | 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 | 3272 |   shows "M[x\<turnstile>n>y]{y:=<c>.P} = Cut <c>.P (x).(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 | 3273 | 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 | 3274 | 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 | 3275 | 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 | 3276 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3278 | 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 | 3279 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3281 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3283 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3285 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3287 | 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 | 3288 | 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 | 3289 | 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 | 3290 | 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 | 3291 | 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 | 3292 | 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 | 3293 | 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 | 3294 | 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 | 3295 | 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 | 3296 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3298 | 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 | 3299 | 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 | 3300 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3302 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3304 | 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 | 3305 | 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 | 3306 | 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 | 3307 | 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 | 3308 | 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 | 3309 | 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 | 3310 | 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 | 3311 | 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 | 3312 | 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 | 3313 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3315 | 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 | 3316 | 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 | 3317 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3318 | apply(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 | 3319 | 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 | 3320 | 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 | 3321 | 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 | 3322 | 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 | 3323 | 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 | 3324 | 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 | 3325 | 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 | 3326 | 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 | 3327 | 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 | 3328 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3329 | apply(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 | 3330 | 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 | 3331 | 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 | 3332 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3333 | apply(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 | 3334 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3335 | apply(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 | 3336 | 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 | 3337 | 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 | 3338 | 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 | 3339 | 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 | 3340 | 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 | 3341 | 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 | 3342 | 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 | 3343 | 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 | 3344 | 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 | 3345 | 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 | 3346 | 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 | 3347 | 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 | 3348 | 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 | 3349 | 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 | 3350 | 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 | 3351 | 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 | 3352 | 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 | 3353 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3354 | apply(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 | 3355 | 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 | 3356 | 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 | 3357 | 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 | 3358 | 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 | 3359 | 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 | 3360 | 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 | 3361 | 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 | 3362 | 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 | 3363 | 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 | 3364 | 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 | 3365 | 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 | 3366 | 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 | 3367 | 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 | 3368 | 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 | 3369 | 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 | 3370 | 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 | 3371 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3374 | 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 | 3375 | 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 | 3376 | [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 | 3377 | | [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 | 3378 | | [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 | 3379 | | [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 | 3380 | | [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 | 3381 | | [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 | 3382 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3384 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3386 | 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 | 3387 | 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 | 3388 | 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 | 3389 | apply(erule_tac 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 | 3390 | apply(auto 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 | 3391 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3394 | 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 | 3395 | 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 | 3396 | 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 | 3397 | apply(erule_tac 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 | 3398 | apply(auto 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 | 3399 | 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 | 3400 | 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 | 3401 | 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 | 3402 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3405 | 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 | 3406 | 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 | 3407 | 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 | 3408 | apply(erule_tac 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 | 3409 | apply(auto 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 | 3410 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3413 | 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 | 3414 | 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 | 3415 | 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 | 3416 | apply(erule_tac 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 | 3417 | apply(auto 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 | 3418 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3421 | 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 | 3422 | 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 | 3423 | 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 | 3424 | apply(erule_tac 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 | 3425 | apply(auto 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 | 3426 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3429 | 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 | 3430 | 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 | 3431 | 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 | 3432 | apply(erule_tac 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 | 3433 | apply(auto 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 | 3434 | 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 | 3435 | 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 | 3436 | 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 | 3437 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3440 | 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 | 3441 | 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 | 3442 | 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 | 3443 | 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 | 3444 | 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 | 3445 | 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 | 3446 | 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 | 3447 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3449 | 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 | 3450 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3452 | 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 | 3453 | by (induct rule: fic.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 | 3454 | (auto simp add: calc_atm simp add: 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 | 3455 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3457 | 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 | 3458 |   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 | 3459 | 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 | 3460 | 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 | 3461 | 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 | 3462 | 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 | 3463 | 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 | 3464 | 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 | 3465 | 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 | 3466 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3468 | 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 | 3469 | 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 | 3470 | 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 | 3471 | 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 | 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(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 | 3474 | 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 | 3475 | 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 | 3476 | 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 | 3477 | 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 | 3478 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3480 | 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 | 3481 | 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 | 3482 | 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 | 3483 | 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 | 3484 | 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 | 3485 | 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 | 3486 | 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 | 3487 | 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 | 3488 | 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 | 3489 | 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 | 3490 | 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 | 3491 | 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 | 3492 | 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 | 3493 | 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 | 3494 | 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 | 3495 | 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 | 3496 | 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 | 3497 | 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 | 3498 | 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 | 3499 | 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 | 3500 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 3502 | 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 | 3503 | 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 | 3504 | 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 | 3505 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3506 | apply(simp 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 | 3507 | 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 | 3508 | 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 | 3509 | 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 | 3510 | 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 | 3511 | 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 | 3512 | 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 | 3513 | 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 | 3514 | 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 | 3515 | 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 | 3516 | 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 | 3517 | 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 | 3518 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3519 | apply(simp add: 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 | 3520 | 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 | 3521 | 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 | 3522 | 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 | 3523 | 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 | 3524 | 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 | 3525 | 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 | 3526 | 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 | 3527 | 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 | 3528 | 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 | 3529 | 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 | 3530 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3533 | 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 | 3534 |   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 | 3535 | 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 | 3536 | 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 | 3537 | 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 | 3538 | 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 | 3539 | 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 | 3540 | 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 | 3541 | 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 | 3542 | 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 | 3543 | 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 | 3544 | 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 | 3545 | 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 | 3546 | 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 | 3547 | 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 | 3548 | 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 | 3549 | 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 | 3550 | 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 | 3551 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3553 | 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 | 3554 | 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 | 3555 | 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 | 3556 | 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 | 3557 | 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 | 3558 | 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 | 3559 | 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 | 3560 | 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 | 3561 | 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 | 3562 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3564 | 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 | 3565 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3567 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3568 | apply(simp add: 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 | 3569 | 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 | 3570 | 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 | 3571 | 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 | 3572 | 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 | 3573 | 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 | 3574 | 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 | 3575 | 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 | 3576 | 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 | 3577 | 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 | 3578 | 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 | 3579 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3581 | 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 | 3582 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3584 | 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 | 3585 | 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 | 3586 | 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 | 3587 | 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 | 3588 | 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 | 3589 | 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 | 3590 | 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 | 3591 | 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 | 3592 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3594 | 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 | 3595 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3597 | 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 | 3598 | 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 | 3599 | 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 | 3600 | 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 | 3601 | 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 | 3602 | 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 | 3603 | 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 | 3604 | 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 | 3605 | 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 | 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(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 | 3608 | 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 | 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(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 | 3611 | 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 | 3612 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3615 | 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 | 3616 |   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 | 3617 | 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 | 3618 | 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 | 3619 | 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 | 3620 | 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 | 3621 | 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 | 3622 | 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 | 3623 | 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 | 3624 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 3626 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3628 | 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 | 3629 | 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 | 3630 | 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 | 3631 | 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 | 3632 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3634 | 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 | 3635 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3637 | 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 | 3638 | 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 | 3639 | 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 | 3640 | 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 | 3641 | 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 | 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 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 | 3644 | 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 | 3645 | 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 | 3646 | 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 | 3647 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3649 | 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 | 3650 | 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 | 3651 | 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 | 3652 | 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 | 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 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 | 3655 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3658 | 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 | 3659 |   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 | 3660 | 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 | 3661 | 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 | 3662 | 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 | 3663 | 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 | 3664 | 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 | 3665 | 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 | 3666 | 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 | 3667 | 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 | 3668 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3670 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3672 | 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 | 3673 | 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 | 3674 | 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 | 3675 | 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 | 3676 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3678 | 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 | 3679 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3680 | 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 | 3681 | 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 | 3682 | 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 | 3683 | 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 | 3684 | 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 | 3685 | 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 | 3686 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3688 | 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 | 3689 | 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 | 3690 | 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 | 3691 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3693 | 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 | 3694 | 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 | 3695 | 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 | 3696 | 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 | 3697 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3698 | 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 | 3699 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3702 | 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 | 3703 |   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 | 3704 | 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 | 3705 | 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 | 3706 | 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 | 3707 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3708 | apply(simp add: 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 | 3709 | 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 | 3710 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3711 | apply(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 | 3712 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3713 | apply(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 | 3714 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3716 | 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 | 3717 | 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 | 3718 | 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 | 3719 | 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 | 3720 | 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 | 3721 | 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 | 3722 | 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 | 3723 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 3725 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3726 | apply(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 | 3727 | 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 | 3728 | 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 | 3729 | 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 | 3730 | 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 | 3731 | 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 | 3732 | 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 | 3733 | 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 | 3734 | 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 | 3735 | 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 | 3736 | 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 | 3737 | 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 | 3738 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3739 | 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 | 3740 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3741 | apply(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 | 3742 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3743 | apply(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 | 3744 | 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 | 3745 | 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 | 3746 | 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 | 3747 | 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 | 3748 | 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 | 3749 | 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 | 3750 | 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 | 3751 | 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 | 3752 | 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 | 3753 | 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 | 3754 | 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 | 3755 | 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 | 3756 | 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 | 3757 | 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 | 3758 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3759 | apply(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 | 3760 | 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 | 3761 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3762 | apply(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 | 3763 | 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 | 3764 | 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 | 3765 | 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 | 3766 | 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 | 3767 | 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 | 3768 | 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 | 3769 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3773 |   l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>l _" [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 | 3774 | 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 | 3775 | LAxR: "\<lbrakk>x\<sharp>M; a\<sharp>b; fic M a\<rbrakk> \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^isub>l 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 | 3776 | | LAxL: "\<lbrakk>a\<sharp>M; x\<sharp>y; fin M x\<rbrakk> \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^isub>l 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 | 3777 | | 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> | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3778 | Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^isub>l Cut <b>.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 | 3779 | | 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> | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>l Cut <a1>.M1 (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 | 3781 | | 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> | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL2 (x).N y) \<longrightarrow>\<^isub>l Cut <a2>.M2 (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 | 3783 | | 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> | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x1).N1" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | | 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> | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x2).N2" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3787 | | 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 | 3788 | c\<sharp>(P,[a].M,b,a); a\<sharp>([c].N,P); y\<sharp>(N,[x].M)\<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 | 3789 | Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>l 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 | 3790 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3792 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3793 | 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 | 3794 | 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 | 3795 | 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 | 3796 | shows "(pi1\<bullet>M) \<longrightarrow>\<^isub>l (pi1\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^isub>l 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 | 3797 | and "(pi2\<bullet>M) \<longrightarrow>\<^isub>l (pi2\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^isub>l 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 | 3798 | 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 | 3799 | 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 | 3800 | 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 | 3801 | 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 | 3802 | 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 | 3803 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3804 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3806 | 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 | 3807 | 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 | 3808 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3809 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3811 | 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 | 3812 | 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 | 3813 | shows "M \<longrightarrow>\<^isub>l M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> 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 | 3814 | and "M \<longrightarrow>\<^isub>l M' \<Longrightarrow> a\<sharp>M \<Longrightarrow> 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 | 3815 | 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 | 3816 | apply(induct rule: l_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 | 3817 | apply(auto simp add: abs_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 | 3818 | 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 | 3819 | 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 | 3820 | 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 | 3821 | 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 | 3822 | apply(induct rule: l_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 | 3823 | apply(auto simp add: abs_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 | 3824 | 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 | 3825 | 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 | 3826 | 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 | 3827 | 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 | 3828 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_LAxR_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 | 3831 | shows "fic M a \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^isub>l 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 | 3832 | 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 | 3833 | 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 | 3834 | 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 | 3835 | 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 | 3836 | have "Cut <a>.M (x).(Ax x b) = Cut <a'>.([(a',a)]\<bullet>M) (x').(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 | 3837 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3838 | also have "\<dots> \<longrightarrow>\<^isub>l ([(a',a)]\<bullet>M)[a'\<turnstile>c>b]" using fs1 fs2 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 | 3839 | 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 | 3840 | 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 | 3841 | 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 | 3842 | 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 | 3843 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_LAxL_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 | 3845 | shows "fin M x \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^isub>l 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 | 3846 | 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 | 3847 | 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 | 3848 | 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 | 3849 | 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 | 3850 | have "Cut <a>.(Ax y a) (x).M = Cut <a'>.(Ax y a') (x').([(x',x)]\<bullet>M)" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3851 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3852 | also have "\<dots> \<longrightarrow>\<^isub>l ([(x',x)]\<bullet>M)[x'\<turnstile>n>y]" using fs1 fs2 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 | 3853 | 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 | 3854 | 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 | 3855 | 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 | 3856 | 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 | 3857 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_LNot_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 | 3859 | shows "\<lbrakk>y\<sharp>N; a\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^isub>l Cut <b>.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 | 3860 | 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 | 3861 | 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 | 3862 | 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 | 3863 | 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 | 3864 | 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 | 3865 | 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 | 3866 | 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 | 3867 | = 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 | 3868 | using f1 f2 f3 f4 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3870 | 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 | 3871 | 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 | 3872 | 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 | 3873 | using f1 f2 f3 f4 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3875 | also have "\<dots> \<longrightarrow>\<^isub>l Cut <b'>.([(b',b)]\<bullet>N) (x').([(x',x)]\<bullet>M)" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3876 | 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 | 3877 | 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 | 3878 | also have "\<dots> = Cut <b>.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 | 3879 | using f1 f2 f3 f4 by (auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3880 | 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 | 3881 | 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 | 3882 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3884 | shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3885 | \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>l Cut <b1>.M1 (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 | 3886 | 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 | 3887 | 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 | 3888 | 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 | 3889 | 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 | 3890 | 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 | 3891 | 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 | 3892 | 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 | 3893 | 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 | 3894 | = 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 | 3895 | 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 | 3896 | 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 | 3897 | 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 | 3898 | apply(auto 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 | 3899 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3900 | 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 | 3901 | (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 | 3902 | 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 | 3903 | 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 | 3904 | 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 | 3905 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> \<longrightarrow>\<^isub>l Cut <b1'>.([(b1',b1)]\<bullet>M1) (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 | 3907 | 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 | 3908 | 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 | 3909 | apply(rule l_redu.intros) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3910 | apply(auto simp add: abs_fresh 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 | 3911 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> = Cut <b1>.M1 (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 | 3913 | using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3914 | 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 | 3915 | 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 | 3916 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3918 | shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3919 | \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).N y) \<longrightarrow>\<^isub>l Cut <b2>.M2 (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 | 3920 | 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 | 3921 | 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 | 3922 | 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 | 3923 | 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 | 3924 | 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 | 3925 | 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 | 3926 | 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 | 3927 | 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 | 3928 | = 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 | 3929 | 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 | 3930 | 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 | 3931 | 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 | 3932 | apply(auto 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 | 3933 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3935 | (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 | 3936 | 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 | 3937 | 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 | 3938 | 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 | 3939 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3940 | also have "\<dots> \<longrightarrow>\<^isub>l Cut <b2'>.([(b2',b2)]\<bullet>M2) (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 | 3941 | 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 | 3942 | 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 | 3943 | apply(rule l_redu.intros) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3944 | apply(auto simp add: abs_fresh 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 | 3945 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> = Cut <b2>.M2 (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 | 3947 | using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3948 | 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 | 3949 | 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 | 3950 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3952 | shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3953 | \<Longrightarrow> Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x1).N1" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3954 | 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 | 3955 | 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 | 3956 | 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 | 3957 | 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 | 3958 | 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 | 3959 | 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 | 3960 | 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 | 3961 | 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 | 3962 | = 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 | 3963 | 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 | 3964 | 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 | 3965 | 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 | 3966 | apply(auto 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 | 3967 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3968 | 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 | 3969 | (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 | 3970 | 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 | 3971 | 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 | 3972 | 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 | 3973 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.([(a',a)]\<bullet>M) (x1').([(x1',x1)]\<bullet>N1)" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3976 | 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 | 3977 | apply(rule l_redu.intros) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3978 | apply(auto simp add: abs_fresh 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 | 3979 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> = Cut <a>.M (x1).N1" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 3982 | 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 | 3983 | 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 | 3984 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 3986 | shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3987 | \<Longrightarrow> Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x2).N2" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 3988 | 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 | 3989 | 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 | 3990 | 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 | 3991 | 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 | 3992 | 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 | 3993 | 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 | 3994 | 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 | 3995 | 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 | 3996 | = 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 | 3997 | 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 | 3998 | 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 | 3999 | 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 | 4000 | apply(auto 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 | 4001 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4003 | (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 | 4004 | 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 | 4005 | 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 | 4006 | 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 | 4007 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4008 | also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.([(a',a)]\<bullet>M) (x2').([(x2',x2)]\<bullet>N2)" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4009 | 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 | 4010 | 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 | 4011 | apply(rule l_redu.intros) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4012 | apply(auto simp add: abs_fresh 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 | 4013 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> = Cut <a>.M (x2).N2" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4015 | using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4016 | 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 | 4017 | 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 | 4018 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4019 | 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 | 4020 | shows "\<lbrakk>z\<sharp>(N,[y].P); b\<sharp>[a].M; a\<sharp>N\<rbrakk> | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4021 | \<Longrightarrow> Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>l 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 | 4022 | 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 | 4023 | 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 | 4024 | 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 | 4025 | 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 | 4026 | 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 | 4027 | 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 | 4028 | 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 | 4029 | 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 | 4030 | 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 | 4031 | = 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 | 4032 | 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 | 4033 | 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 | 4034 | 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 | 4035 | apply(auto 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 | 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 | 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 | 4038 | (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 | 4039 | 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 | 4040 | 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 | 4041 | 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 | 4042 | 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 | 4043 | 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 | 4044 | 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 | 4045 | 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 | 4046 | 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 | 4047 | 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 | 4048 | apply(rule conjI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4049 | apply(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 | 4050 | 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 | 4051 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4052 | also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.(Cut <c'>.([(c',c)]\<bullet>N) (x').([(a',a)]\<bullet>[(x',x)]\<bullet>M)) (y').([(y',y)]\<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 | 4053 | 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 | 4054 | 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 | 4055 | apply(rule l_redu.intros) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4056 | apply(auto simp add: abs_fresh 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 | 4057 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4058 | 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 | 4059 | 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 | 4060 | 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 | 4061 | 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 | 4062 | 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 | 4063 | 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 | 4064 | 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 | 4065 | 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 | 4066 | 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 | 4067 | 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 | 4068 | 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 | 4069 | apply(auto simp add: fresh_prod fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4070 | 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 | 4071 | 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 | 4072 | 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 | 4073 | 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 | 4074 | 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 | 4075 | 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 | 4076 | 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 | 4077 | 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 | 4078 | 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 | 4079 | 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 | 4080 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4082 | 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 | 4083 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4084 | 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 | 4085 | 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 | 4086 | 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 | 4087 | 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 | 4088 | 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 | 4089 | 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 | 4090 | apply(auto simp add: alpha_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 | 4091 | 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 | 4092 | 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 | 4093 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4094 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4096 | 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 | 4097 | 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 | 4098 | 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 | 4099 | 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 | 4100 | 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 | 4101 | apply(auto simp add: alpha_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 | 4102 | 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 | 4103 | 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 | 4104 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4106 | 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 | 4107 | 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 | 4108 | 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 | 4109 | 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 | 4110 | 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 | 4111 | 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 | 4112 | 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 | 4113 | apply(auto simp add: alpha_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 | 4114 | 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 | 4115 | 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 | 4116 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4118 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4120 | lemma Cut_l_redu_elim: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4121 | assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^isub>l R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4122 | 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 | 4123 | (\<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 | 4124 | (\<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 | 4125 | \<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 | 4126 | (\<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 | 4127 | \<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 | 4128 | (\<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 | 4129 | \<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 | 4130 | (\<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 | 4131 | \<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 | 4132 | (\<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 | 4133 | 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 | 4134 | 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 | 4135 | 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 | 4136 | 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 | 4137 | (* 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 | 4138 | 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 | 4139 | 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 | 4140 | 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 | 4141 | 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 | 4142 | 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 | 4143 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4144 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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: 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 | 4146 | 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 | 4147 | 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 | 4148 | (* 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 | 4149 | 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 | 4150 | 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 | 4151 | 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 | 4152 | 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 | 4153 | 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 | 4154 | 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 | 4155 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4156 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4157 | apply(simp 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 | 4158 | 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 | 4159 | 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 | 4160 | 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 | 4161 | (* 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 | 4162 | 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 | 4163 | 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 | 4164 | 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 | 4165 | 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 | 4166 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4167 | apply(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 | 4168 | 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 | 4169 | 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 | 4170 | 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 | 4171 | 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 | 4172 | 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 | 4173 | 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 | 4174 | 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 | 4175 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4177 | 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 | 4178 | 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 | 4179 | 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 | 4180 | 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 | 4181 | 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 | 4182 | apply(auto simp add: calc_atm abs_fresh fresh_left)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4183 | apply(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 | 4184 | 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 | 4185 | 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 | 4186 | 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 | 4187 | 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 | 4188 | 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 | 4189 | (* 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 | 4190 | 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 | 4191 | 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 | 4192 | 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 | 4193 | 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 | 4194 | 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 | 4195 | 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 | 4196 | 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 | 4197 | 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 | 4198 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4199 | 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 | 4200 | 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 | 4201 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4203 | 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 | 4204 | 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 | 4205 | 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 | 4206 | 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 | 4207 | 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 | 4208 | 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 | 4209 | 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 | 4210 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4211 | apply(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 | 4212 | 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 | 4213 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4215 | 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 | 4216 | 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 | 4217 | 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 | 4218 | 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 | 4219 | apply(auto simp add: fresh_left calc_atm abs_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4220 | 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 | 4221 | 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 | 4222 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4223 | 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 | 4224 | 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 | 4225 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4227 | 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 | 4228 | 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 | 4229 | 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 | 4230 | 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 | 4231 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4232 | 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 | 4233 | 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 | 4234 | 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 | 4235 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4237 | 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 | 4238 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4239 | apply(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 | 4240 | 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 | 4241 | 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 | 4242 | 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 | 4243 | 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 | 4244 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4245 | 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 | 4246 | 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 | 4247 | 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 | 4248 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4249 | apply(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 | 4250 | 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 | 4251 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4253 | 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 | 4254 | 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 | 4255 | 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 | 4256 | 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 | 4257 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4258 | 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 | 4259 | (* 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 | 4260 | 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 | 4261 | 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 | 4262 | 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 | 4263 | 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 | 4264 | 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 | 4265 | 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 | 4266 | 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 | 4267 | 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 | 4268 | 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 | 4269 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4270 | apply(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 | 4271 | 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 | 4272 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4274 | 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 | 4275 | 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 | 4276 | 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 | 4277 | 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 | 4278 | 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 | 4279 | 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 | 4280 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4281 | apply(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 | 4282 | 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 | 4283 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4285 | 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 | 4286 | 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 | 4287 | 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 | 4288 | 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 | 4289 | apply(auto simp add: fresh_left calc_atm abs_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4290 | 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 | 4291 | 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 | 4292 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4293 | apply(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 | 4294 | 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 | 4295 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4296 | apply(rule 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 | 4297 | 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 | 4298 | 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 | 4299 | 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 | 4300 | 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 | 4301 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4302 | 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 | 4303 | 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 | 4304 | 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 | 4305 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4307 | 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 | 4308 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4309 | apply(rule 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 | 4310 | 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 | 4311 | 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 | 4312 | 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 | 4313 | 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 | 4314 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4315 | 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 | 4316 | 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 | 4317 | 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 | 4318 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4320 | 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 | 4321 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4322 | apply(rule 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 | 4323 | 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 | 4324 | 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 | 4325 | 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 | 4326 | 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 | 4327 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4328 | 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 | 4329 | (* 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 | 4330 | 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 | 4331 | 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 | 4332 | 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 | 4333 | 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 | 4334 | 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 | 4335 | 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 | 4336 | 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 | 4337 | 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 | 4338 | 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 | 4339 | 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 | 4340 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4342 | 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 | 4343 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4344 | apply(rule 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 | 4345 | 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 | 4346 | 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 | 4347 | 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 | 4348 | 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 | 4349 | 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 | 4350 | 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 | 4351 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4352 | apply(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 | 4353 | 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 | 4354 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4356 | 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 | 4357 | 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 | 4358 | 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 | 4359 | 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 | 4360 | 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 | 4361 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4362 | 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 | 4363 | 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 | 4364 | 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 | 4365 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4366 | 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 | 4367 | 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 | 4368 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4370 | 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 | 4371 | 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 | 4372 | 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 | 4373 | 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 | 4374 | 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 | 4375 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4376 | 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 | 4377 | (* 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 | 4378 | 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 | 4379 | 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 | 4380 | 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 | 4381 | 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 | 4382 | 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 | 4383 | 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 | 4384 | 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 | 4385 | 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 | 4386 | 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 | 4387 | 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 | 4388 | 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 | 4389 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4390 | apply(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 | 4391 | 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 | 4392 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4394 | 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 | 4395 | 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 | 4396 | 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 | 4397 | 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 | 4398 | 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 | 4399 | 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 | 4400 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4401 | apply(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 | 4402 | 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 | 4403 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4404 | apply(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 | 4405 | 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 | 4406 | 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 | 4407 | 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 | 4408 | 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 | 4409 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4410 | 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 | 4411 | 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 | 4412 | 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 | 4413 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 4415 | 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 | 4416 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4417 | apply(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 | 4418 | 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 | 4419 | 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 | 4420 | 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 | 4421 | 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 | 4422 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4423 | 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 | 4424 | (* 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 | 4425 | 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 | 4426 | 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 | 4427 | 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 | 4428 | 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 | 4429 | 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 | 4430 | 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 | 4431 | 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 | 4432 | 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 | 4433 | 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 | 4434 | 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 | 4435 | 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 | 4436 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4437 | apply(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 | 4438 | 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 | 4439 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4441 | 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 | 4442 | 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 | 4443 | 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 | 4444 | 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 | 4445 | apply(generate_fresh "name") | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4446 | apply(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 | 4447 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4448 | apply(drule_tac 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 | 4449 | 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 | 4450 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4451 | apply(rule 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 | 4452 | 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 | 4453 | 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 | 4454 | 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 | 4455 | 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 | 4456 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4457 | 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 | 4458 | 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 | 4459 | 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 | 4460 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4461 | apply(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 | 4462 | 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 | 4463 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4464 | apply(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 | 4465 | 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 | 4466 | 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 | 4467 | 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 | 4468 | 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 | 4469 | apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4470 | 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 | 4471 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4473 | 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 | 4474 |   c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>c _" [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 | 4475 | 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 | 4476 |   left[intro]:  "\<lbrakk>\<not>fic M a; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c M{a:=(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 | 4477 | | right[intro]: "\<lbrakk>\<not>fin N x; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c N{x:=<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 | 4478 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4480 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4482 | 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 | 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 better_left[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 | 4485 |   shows "\<not>fic M a \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c M{a:=(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 | 4486 | 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 | 4487 | 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 | 4488 | 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 | 4489 | 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 | 4490 | have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (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 | 4491 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4492 |   also have "\<dots> \<longrightarrow>\<^isub>c ([(a',a)]\<bullet>M){a':=(x').([(x',x)]\<bullet>N)}" using fs1 fs2 not_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 | 4493 | 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 | 4494 | 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 | 4495 | 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 | 4496 | 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 | 4497 | 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 | 4498 | 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 | 4499 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   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 | 4501 | 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 | 4502 | 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 | 4503 | 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 | 4504 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4505 | lemma better_right[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 | 4506 |   shows "\<not>fin N x \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c N{x:=<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 | 4507 | 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 | 4508 | 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 | 4509 | 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 | 4510 | 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 | 4511 | have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (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 | 4512 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4513 |   also have "\<dots> \<longrightarrow>\<^isub>c ([(x',x)]\<bullet>N){x':=<a'>.([(a',a)]\<bullet>M)}" using fs1 fs2 not_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 | 4514 | 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 | 4515 | 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 | 4516 | 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 | 4517 | 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 | 4518 | 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 | 4519 | 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 | 4520 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4521 |   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 | 4522 | 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 | 4523 | 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 | 4524 | 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 | 4525 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4527 | 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 | 4528 | 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 | 4529 | shows "M \<longrightarrow>\<^isub>c M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> 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 | 4530 | and "M \<longrightarrow>\<^isub>c M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> 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 | 4531 | 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 | 4532 | apply(induct rule: c_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 | 4533 | apply(auto simp add: abs_fresh rename_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 | 4534 | apply(induct rule: c_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 | 4535 | apply(auto simp add: abs_fresh rename_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 | 4536 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4537 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4539 |   a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>a _" [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 | 4540 | 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 | 4541 | al_redu[intro]: "M\<longrightarrow>\<^isub>l M' \<Longrightarrow> M \<longrightarrow>\<^isub>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 | 4542 | | ac_redu[intro]: "M\<longrightarrow>\<^isub>c M' \<Longrightarrow> M \<longrightarrow>\<^isub>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 | 4543 | | a_Cut_l: "\<lbrakk>a\<sharp>N; x\<sharp>M; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M' (x).N" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4544 | | a_Cut_r: "\<lbrakk>a\<sharp>N; x\<sharp>M; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M (x).N'" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4545 | | a_NotL[intro]: "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^isub>a 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 | 4546 | | a_NotR[intro]: "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^isub>a NotR (x).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 | 4547 | | a_AndR_l: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M' <b>.N c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4548 | | a_AndR_r: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M <b>.N' c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4549 | | a_AndL1: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a AndL1 (x).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 | 4550 | | a_AndL2: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a AndL2 (x).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 | 4551 | | a_OrL_l: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M' (y).N z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4552 | | a_OrL_r: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M (y).N' z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4553 | | a_OrR1: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a OrR1 <a>.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 | 4554 | | a_OrR2: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a OrR2 <a>.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 | 4555 | | a_ImpL_l: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M' (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 | 4556 | | a_ImpL_r: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M (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 | 4557 | | a_ImpR: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a ImpR (x).<a>.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 | 4558 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4560 | 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 | 4561 | 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 | 4562 | shows "M \<longrightarrow>\<^isub>a M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> 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 | 4563 | and "M \<longrightarrow>\<^isub>a M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> 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 | 4564 | 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 | 4565 | 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 | 4566 | 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 | 4567 | apply(simp add: 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 | 4568 | apply(auto simp add: abs_fresh abs_supp fin_supp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4569 | 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 | 4570 | 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 | 4571 | apply(simp add: 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 | 4572 | apply(auto simp add: abs_fresh abs_supp fin_supp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4573 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4574 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4575 | 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 | 4576 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4578 | 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 | 4579 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_CutL_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 | 4581 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M' (x).N" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4582 | 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 | 4583 | assume red: "M\<longrightarrow>\<^isub>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 | 4584 | 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 | 4585 | 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 | 4586 | have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (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 | 4587 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4588 | also have "\<dots> \<longrightarrow>\<^isub>a Cut <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N)" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4590 | also have "\<dots> = Cut <a>.M' (x).N" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4591 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4592 | 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 | 4593 | 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 | 4594 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4595 | lemma better_CutR_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 | 4596 | shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M (x).N'" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4597 | 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 | 4598 | assume red: "N\<longrightarrow>\<^isub>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 | 4599 | 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 | 4600 | 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 | 4601 | have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (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 | 4602 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4603 | also have "\<dots> \<longrightarrow>\<^isub>a Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N')" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4604 | 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 | 4605 | also have "\<dots> = Cut <a>.M (x).N'" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4606 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4607 | 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 | 4608 | 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 | 4609 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4610 | lemma better_AndRL_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 | 4611 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M' <b>.N c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4612 | 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 | 4613 | assume red: "M\<longrightarrow>\<^isub>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 | 4614 | 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 | 4615 | 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 | 4616 | have "AndR <a>.M <b>.N c = AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4617 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4618 | also have "\<dots> \<longrightarrow>\<^isub>a AndR <a'>.([(a',a)]\<bullet>M') <b'>.([(b',b)]\<bullet>N) c" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4620 | also have "\<dots> = AndR <a>.M' <b>.N c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4621 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4622 | 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 | 4623 | 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 | 4624 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4625 | lemma better_AndRR_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 | 4626 | shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M <b>.N' c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4627 | 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 | 4628 | assume red: "N\<longrightarrow>\<^isub>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 | 4629 | 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 | 4630 | 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 | 4631 | have "AndR <a>.M <b>.N c = AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4632 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4633 | also have "\<dots> \<longrightarrow>\<^isub>a AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N') c" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4634 | 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 | 4635 | also have "\<dots> = AndR <a>.M <b>.N' c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4636 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4637 | 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 | 4638 | 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 | 4639 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4640 | lemma better_AndL1_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 | 4641 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a AndL1 (x).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 | 4642 | 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 | 4643 | assume red: "M\<longrightarrow>\<^isub>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 | 4644 | 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 | 4645 | have "AndL1 (x).M y = AndL1 (x').([(x',x)]\<bullet>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 | 4646 | using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4647 | also have "\<dots> \<longrightarrow>\<^isub>a AndL1 (x').([(x',x)]\<bullet>M') y" using fs1 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4649 | also have "\<dots> = AndL1 (x).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 | 4650 | using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4651 | 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 | 4652 | 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 | 4653 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4654 | lemma better_AndL2_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 | 4655 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a AndL2 (x).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 | 4656 | 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 | 4657 | assume red: "M\<longrightarrow>\<^isub>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 | 4658 | 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 | 4659 | have "AndL2 (x).M y = AndL2 (x').([(x',x)]\<bullet>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 | 4660 | using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4661 | also have "\<dots> \<longrightarrow>\<^isub>a AndL2 (x').([(x',x)]\<bullet>M') y" using fs1 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4663 | also have "\<dots> = AndL2 (x).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 | 4664 | using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4665 | 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 | 4666 | 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 | 4667 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_OrLL_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 | 4669 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M' (y).N z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4670 | 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 | 4671 | assume red: "M\<longrightarrow>\<^isub>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 | 4672 | 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 | 4673 | 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 | 4674 | have "OrL (x).M (y).N z = OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>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 | 4675 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4676 | also have "\<dots> \<longrightarrow>\<^isub>a OrL (x').([(x',x)]\<bullet>M') (y').([(y',y)]\<bullet>N) z" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4678 | also have "\<dots> = OrL (x).M' (y).N z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4679 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4680 | 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 | 4681 | 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 | 4682 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_OrLR_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 | 4684 | shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M (y).N' z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4685 | 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 | 4686 | assume red: "N\<longrightarrow>\<^isub>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 | 4687 | 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 | 4688 | 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 | 4689 | have "OrL (x).M (y).N z = OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>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 | 4690 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4691 | also have "\<dots> \<longrightarrow>\<^isub>a OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N') z" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4693 | also have "\<dots> = OrL (x).M (y).N' z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4694 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4695 | 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 | 4696 | 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 | 4697 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_OrR1_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 | 4699 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a OrR1 <a>.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 | 4700 | 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 | 4701 | assume red: "M\<longrightarrow>\<^isub>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 | 4702 | 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 | 4703 | have "OrR1 <a>.M 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 | 4704 | using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4705 | also have "\<dots> \<longrightarrow>\<^isub>a OrR1 <a'>.([(a',a)]\<bullet>M') b" using fs1 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4707 | also have "\<dots> = OrR1 <a>.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 | 4708 | using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4709 | 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 | 4710 | 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 | 4711 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_OrR2_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 | 4713 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a OrR2 <a>.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 | 4714 | 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 | 4715 | assume red: "M\<longrightarrow>\<^isub>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 | 4716 | 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 | 4717 | have "OrR2 <a>.M 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 | 4718 | using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4719 | also have "\<dots> \<longrightarrow>\<^isub>a OrR2 <a'>.([(a',a)]\<bullet>M') b" using fs1 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4721 | also have "\<dots> = OrR2 <a>.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 | 4722 | using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4723 | 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 | 4724 | 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 | 4725 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_ImpLL_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 | 4727 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M' (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 | 4728 | 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 | 4729 | assume red: "M\<longrightarrow>\<^isub>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 | 4730 | 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 | 4731 | 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 | 4732 | have "ImpL <a>.M (x).N y = ImpL <a'>.([(a',a)]\<bullet>M) (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 | 4733 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4734 | also have "\<dots> \<longrightarrow>\<^isub>a ImpL <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N) y" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4736 | also have "\<dots> = ImpL <a>.M' (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 | 4737 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4738 | 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 | 4739 | 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 | 4740 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_ImpLR_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 | 4742 | shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M (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 | 4743 | 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 | 4744 | assume red: "N\<longrightarrow>\<^isub>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 | 4745 | 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 | 4746 | 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 | 4747 | have "ImpL <a>.M (x).N y = ImpL <a'>.([(a',a)]\<bullet>M) (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 | 4748 | using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4749 | also have "\<dots> \<longrightarrow>\<^isub>a ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N') y" using fs1 fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4751 | also have "\<dots> = ImpL <a>.M (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 | 4752 | using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4753 | 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 | 4754 | 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 | 4755 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma better_ImpR_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 | 4757 | shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a ImpR (x).<a>.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 | 4758 | 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 | 4759 | assume red: "M\<longrightarrow>\<^isub>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 | 4760 | 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 | 4761 | have "ImpR (x).<a>.M b = ImpR (x).<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 | 4762 | using fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod 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 | 4763 | also have "\<dots> \<longrightarrow>\<^isub>a ImpR (x).<a'>.([(a',a)]\<bullet>M') b" using fs2 red | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4765 | also have "\<dots> = ImpR (x).<a>.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 | 4766 | using fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm 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 | 4767 | 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 | 4768 | 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 | 4769 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* axioms do not 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 | 4771 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4773 | shows "Ax x a \<longrightarrow>\<^isub>l M \<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 | 4774 | 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 | 4775 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4777 | shows "Ax x a \<longrightarrow>\<^isub>c M \<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 | 4778 | 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 | 4779 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4781 | shows "Ax x a \<longrightarrow>\<^isub>a M \<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 | 4782 | apply(erule_tac a_redu.cases) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4783 | apply(auto 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 | 4784 | 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 | 4785 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 4787 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4791 | assumes a: "NotL <a>.M x \<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4792 | shows "\<exists>M'. R = NotL <a>.M' x \<and> M\<longrightarrow>\<^isub>aM'" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4794 | 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 | 4795 | 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 | 4796 | 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 | 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(rotate_tac 1) | 
| 
9be4ab2acc13
split Class.thy into parts to 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_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 | 4800 | 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 | 4801 | 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 | 4802 | apply(auto simp add: alpha 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 | 4803 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4804 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4805 | 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 | 4806 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4807 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4808 | 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 | 4809 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4812 | assumes a: "NotR (x).M a \<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4813 | shows "\<exists>M'. R = NotR (x).M' a \<and> M\<longrightarrow>\<^isub>aM'" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4815 | 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 | 4816 | 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 | 4817 | 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 | 4818 | 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 | 4819 | apply(rotate_tac 1) | 
| 
9be4ab2acc13
split Class.thy into parts to 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_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 | 4821 | 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 | 4822 | 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 | 4823 | apply(auto simp add: alpha 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 | 4824 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4825 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4826 | 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 | 4827 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4828 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4829 | 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 | 4830 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4833 | assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4834 | shows "(\<exists>M'. R = AndR <a>.M' <b>.N c \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = AndR <a>.M <b>.N' c \<and> N\<longrightarrow>\<^isub>aN')" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4836 | 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 | 4837 | 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 | 4838 | 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 | 4839 | 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 | 4840 | 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 | 4841 | 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 | 4842 | 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 | 4843 | 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 | 4844 | apply(auto simp add: alpha a_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 | 4845 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4846 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4847 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4848 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4849 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4850 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4851 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4852 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4853 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4854 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4855 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4856 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4857 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4858 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4859 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4860 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4861 | 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 | 4862 | apply(auto simp add: alpha a_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 | 4863 | apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4864 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4865 | apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4866 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4867 | apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4868 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4869 | apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4870 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4871 | apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4872 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4873 | apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4874 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4875 | apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4876 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4877 | apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4878 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4879 | 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 | 4880 | 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 | 4881 | 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 | 4882 | 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 | 4883 | 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 | 4884 | apply(auto simp add: alpha a_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 | 4885 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4886 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4887 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4888 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4889 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4890 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4891 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4892 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4893 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4894 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4895 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4896 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4897 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4898 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4899 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4900 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4901 | 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 | 4902 | apply(auto simp add: alpha a_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 | 4903 | apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4904 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4905 | apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4906 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4907 | apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4908 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4909 | apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4910 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4911 | apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4912 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4913 | apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4914 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4915 | apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4916 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4917 | apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4918 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4921 | lemma 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 | 4922 | assumes a: "AndL1 (x).M y \<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4923 | shows "\<exists>M'. R = AndL1 (x).M' y \<and> M\<longrightarrow>\<^isub>aM'" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4925 | 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 | 4926 | 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 | 4927 | 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 | 4928 | 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 | 4929 | apply(rotate_tac 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 | 4930 | 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 | 4931 | 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 | 4932 | 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 | 4933 | apply(auto simp add: alpha 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 | 4934 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4935 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4936 | 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 | 4937 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4938 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4939 | 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 | 4940 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4943 | assumes a: "AndL2 (x).M y \<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4944 | shows "\<exists>M'. R = AndL2 (x).M' y \<and> M\<longrightarrow>\<^isub>aM'" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4946 | 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 | 4947 | 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 | 4948 | 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 | 4949 | 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 | 4950 | apply(rotate_tac 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 | 4951 | 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 | 4952 | 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 | 4953 | 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 | 4954 | apply(auto simp add: alpha 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 | 4955 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4956 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4957 | 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 | 4958 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4959 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 4960 | 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 | 4961 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 4964 | assumes a: "OrL (x).M (y).N z\<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4965 | shows "(\<exists>M'. R = OrL (x).M' (y).N z \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = OrL (x).M (y).N' z \<and> N\<longrightarrow>\<^isub>aN')" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 4967 | 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 | 4968 | 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 | 4969 | 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 | 4970 | 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 | 4971 | 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 | 4972 | 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 | 4973 | 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 | 4974 | 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 | 4975 | apply(auto simp add: alpha a_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 | 4976 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4977 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4978 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4979 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4980 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4981 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4982 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4983 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4984 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4985 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4986 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4987 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4988 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4989 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4990 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4991 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4992 | 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 | 4993 | apply(auto simp add: alpha a_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 | 4994 | apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4995 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4996 | apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4997 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4998 | apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 4999 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5000 | apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5001 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5002 | apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5003 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5004 | apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5005 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5006 | apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5007 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5008 | apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5009 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5010 | 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 | 5011 | 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 | 5012 | 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 | 5013 | 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 | 5014 | 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 | 5015 | apply(auto simp add: alpha a_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 | 5016 | apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5017 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5018 | apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5019 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5020 | apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5021 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5022 | apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5023 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5024 | apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5025 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5026 | apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5027 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5028 | apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5029 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5030 | apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5031 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5032 | 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 | 5033 | apply(auto simp add: alpha a_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 | 5034 | apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5035 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5036 | apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5037 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5038 | apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5039 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5040 | apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5041 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5042 | apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5043 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5044 | apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5045 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5046 | apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5047 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5048 | apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5049 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5050 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5052 | lemma 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 | 5053 | assumes a: "OrR1 <a>.M b \<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5054 | shows "\<exists>M'. R = OrR1 <a>.M' b \<and> M\<longrightarrow>\<^isub>aM'" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5056 | 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 | 5057 | 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 | 5058 | 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 | 5059 | 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 | 5060 | apply(rotate_tac 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 | 5061 | 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 | 5062 | 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 | 5063 | 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 | 5064 | apply(auto simp add: alpha 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 | 5065 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5066 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 5067 | 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 | 5068 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5069 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 5070 | 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 | 5071 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 5074 | assumes a: "OrR2 <a>.M b \<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5075 | shows "\<exists>M'. R = OrR2 <a>.M' b \<and> M\<longrightarrow>\<^isub>aM'" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5077 | 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 | 5078 | 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 | 5079 | 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 | 5080 | 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 | 5081 | apply(rotate_tac 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 | 5082 | 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 | 5083 | 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 | 5084 | 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 | 5085 | apply(auto simp add: alpha 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 | 5086 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5087 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 5088 | 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 | 5089 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5090 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm 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 | 5091 | 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 | 5092 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 5095 | assumes a: "ImpL <a>.M (y).N z\<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5096 | shows "(\<exists>M'. R = ImpL <a>.M' (y).N z \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = ImpL <a>.M (y).N' z \<and> N\<longrightarrow>\<^isub>aN')" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5097 | using a | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5099 | 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 | 5100 | 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 | 5101 | 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 | 5102 | 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 | 5103 | 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 | 5104 | 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 | 5105 | 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 | 5106 | apply(auto simp add: alpha a_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 | 5107 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5108 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5109 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5110 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5111 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5112 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5113 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5114 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5115 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5116 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5117 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5118 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5119 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5120 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5121 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5122 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5123 | 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 | 5124 | apply(auto simp add: alpha a_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 | 5125 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5126 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5127 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5128 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5129 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5130 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5131 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5132 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5133 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5134 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5135 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5136 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5137 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5138 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5139 | apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5140 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5141 | 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 | 5142 | 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 | 5143 | 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 | 5144 | 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 | 5145 | 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 | 5146 | apply(auto simp add: alpha a_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 | 5147 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5148 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5149 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5150 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5151 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5152 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5153 | apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5154 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5155 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5156 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5157 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5158 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5159 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5160 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5161 | apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5162 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5163 | 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 | 5164 | apply(auto simp add: alpha a_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 | 5165 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5166 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5167 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5168 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5169 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5170 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5171 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5172 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5173 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5174 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5175 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5176 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5177 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5178 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5179 | apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5180 | apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5181 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5182 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 5184 | assumes a: "ImpR (x).<a>.M b \<longrightarrow>\<^isub>a R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5185 | shows "\<exists>M'. R = ImpR (x).<a>.M' b \<and> M\<longrightarrow>\<^isub>aM'" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5186 | 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 | 5187 | 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 | 5188 | 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 | 5189 | 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 | 5190 | 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 | 5191 | apply(rotate_tac 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 | 5192 | 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 | 5193 | 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 | 5194 | 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 | 5195 | apply(auto simp add: alpha a_redu.eqvt abs_perm 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 | 5196 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5197 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5198 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5199 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5200 | apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5201 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5202 | apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5203 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5204 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5205 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5206 | apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5207 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5208 | apply(rule_tac x="([(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5209 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5210 | 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 | 5211 | 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 | 5212 | 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 | 5213 | 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 | 5214 | apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5215 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5216 | 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 | 5217 | 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 | 5218 | 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 | 5219 | 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 | 5220 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5221 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5222 | apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5223 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5224 | apply(rule_tac x="([(a,aa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5225 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5226 | 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 | 5227 | 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 | 5228 | 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 | 5229 | 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 | 5230 | apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5231 | apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu 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 | 5232 | 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 | 5233 | 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 | 5234 | 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 | 5235 | 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 | 5236 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* Transitive Closure*}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | abbreviation | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5241 |  a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>a* _" [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 | 5242 | 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 | 5243 | "M \<longrightarrow>\<^isub>a* M' \<equiv> (a_redu)^** M 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 | 5244 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 5246 | assumes a: "M \<longrightarrow>\<^isub>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 | 5247 | shows "M \<longrightarrow>\<^isub>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 | 5248 | 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 | 5249 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_starE: | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | assumes a: "M \<longrightarrow>\<^isub>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 | 5252 | shows "M = M' \<or> (\<exists>N. M \<longrightarrow>\<^isub>a N \<and> N \<longrightarrow>\<^isub>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 | 5253 | 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 | 5254 | 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 | 5255 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5256 | lemma a_star_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 | 5257 | shows "M \<longrightarrow>\<^isub>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 | 5258 | 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 | 5259 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_trans[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 | 5261 | assumes a1: "M1\<longrightarrow>\<^isub>a* 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 | 5262 | and a2: "M2\<longrightarrow>\<^isub>a* M3" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | shows "M1 \<longrightarrow>\<^isub>a* M3" | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 5265 | 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 | 5266 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5267 | text {* congruence rules for \<longrightarrow>\<^isub>a* *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma ax_do_not_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 | 5270 | shows "Ax x a \<longrightarrow>\<^isub>a* M \<Longrightarrow> M = Ax x a" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5271 | 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 | 5272 | 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 | 5273 | 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 | 5274 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5275 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_CutL: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5279 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M' (x).N" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5280 | 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 | 5281 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_CutR: | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M (x).N'" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5284 | 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 | 5285 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_Cut: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5287 | "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M' (x).N'" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5288 | 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 | 5289 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_NotR: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5291 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^isub>a* NotR (x).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 | 5292 | 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 | 5293 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_NotL: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5295 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^isub>a* 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 | 5296 | 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 | 5297 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_AndRL: | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M' <b>.N c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5300 | 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 | 5301 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5302 | lemma a_star_AndRR: | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M <b>.N' c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5304 | 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 | 5305 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_AndR: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5307 | "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M' <b>.N' c" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5308 | 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 | 5309 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_AndL1: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5311 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a* AndL1 (x).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 | 5312 | 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 | 5313 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 5315 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a* AndL2 (x).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 | 5316 | 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 | 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 a_star_OrLL: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5319 | "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M' (y).N z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5320 | 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 | 5321 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5322 | lemma a_star_OrLR: | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M (y).N' z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5324 | 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 | 5325 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_OrL: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5327 | "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M' (y).N' z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5328 | 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 | 5329 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_OrR1: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5331 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a* OrR1 <a>.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 | 5332 | 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 | 5333 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_OrR2: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5335 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a* OrR2 <a>.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 | 5336 | 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 | 5337 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5338 | lemma a_star_ImpLL: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5339 | "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M' (y).N z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5340 | 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 | 5341 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_ImpLR: | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M (y).N' z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5344 | 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 | 5345 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_ImpL: | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5347 | "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M' (y).N' z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5348 | 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 | 5349 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma 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 | 5351 | "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a* ImpR (x).<a>.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 | 5352 | 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 | 5353 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5354 | 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 | 5355 | 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 | 5356 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_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 | 5358 | assumes a: "NotL <a>.M x \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5359 | shows "\<exists>M'. R = NotL <a>.M' x \<and> M \<longrightarrow>\<^isub>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 | 5360 | 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 | 5361 | 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 | 5362 | 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 | 5363 | 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 | 5364 | 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 | 5365 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_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 | 5368 | assumes a: "NotR (x).M a \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5369 | shows "\<exists>M'. R = NotR (x).M' a \<and> M \<longrightarrow>\<^isub>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 | 5370 | 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 | 5371 | 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 | 5372 | 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 | 5373 | 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 | 5374 | 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 | 5375 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_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 | 5378 | assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5379 | shows "(\<exists>M' N'. R = AndR <a>.M' <b>.N' c \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>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 | 5380 | 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 | 5381 | 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 | 5382 | 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 | 5383 | 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 | 5384 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5385 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5386 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5387 | lemma a_star_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 | 5388 | assumes a: "AndL1 (x).M y \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5389 | shows "\<exists>M'. R = AndL1 (x).M' y \<and> M \<longrightarrow>\<^isub>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 | 5390 | 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 | 5391 | 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 | 5392 | 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 | 5393 | 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 | 5394 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5395 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_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 | 5398 | assumes a: "AndL2 (x).M y \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5399 | shows "\<exists>M'. R = AndL2 (x).M' y \<and> M \<longrightarrow>\<^isub>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 | 5400 | 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 | 5401 | 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 | 5402 | 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 | 5403 | 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 | 5404 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5405 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_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 | 5408 | assumes a: "OrL (x).M (y).N z \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5409 | shows "(\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>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 | 5410 | 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 | 5411 | 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 | 5412 | 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 | 5413 | 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 | 5414 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5415 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_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 | 5418 | assumes a: "OrR1 <a>.M y \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5419 | shows "\<exists>M'. R = OrR1 <a>.M' y \<and> M \<longrightarrow>\<^isub>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 | 5420 | 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 | 5421 | 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 | 5422 | 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 | 5423 | 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 | 5424 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5425 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma a_star_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 | 5428 | assumes a: "OrR2 <a>.M y \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5429 | shows "\<exists>M'. R = OrR2 <a>.M' y \<and> M \<longrightarrow>\<^isub>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 | 5430 | 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 | 5431 | 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 | 5432 | 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 | 5433 | 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 | 5434 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5435 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5437 | lemma a_star_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 | 5438 | assumes a: "ImpR (x).<a>.M y \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5439 | shows "\<exists>M'. R = ImpR (x).<a>.M' y \<and> M \<longrightarrow>\<^isub>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 | 5440 | 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 | 5441 | 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 | 5442 | 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 | 5443 | 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 | 5444 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5445 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5446 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5447 | lemma a_star_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 | 5448 | assumes a: "ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* R" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5449 | shows "(\<exists>M' N'. R = ImpL <a>.M' (y).N' z \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>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 | 5450 | 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 | 5451 | 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 | 5452 | 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 | 5453 | 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 | 5454 | apply(auto simp add: alpha trm.inject) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5455 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* Substitution *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5458 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5460 |   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 | 5461 | 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 | 5462 | 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 | 5463 | 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 | 5464 | 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 | 5465 | 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 | 5466 | 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 | 5467 | 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 | 5468 | 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 | 5469 | 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 | 5470 | 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 | 5471 | 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 | 5472 | 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 | 5473 | 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 | 5474 | 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 | 5475 | 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 | 5476 | 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 | 5477 | 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 | 5478 | 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 | 5479 | 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 | 5480 | 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 | 5481 | 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 | 5482 | 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 | 5483 | 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 | 5484 | 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 | 5485 | 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 | 5486 | 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 | 5487 | 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 | 5488 | 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 | 5489 | 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 | 5490 | 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 | 5491 | 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 | 5492 | 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 | 5493 | 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 | 5494 | 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 | 5495 | 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 | 5496 | 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 | 5497 | 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 | 5498 | 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 | 5499 | 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 | 5500 | 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 | 5501 | 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 | 5502 | 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 | 5503 | 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 | 5504 | 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 | 5505 | 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 | 5506 | 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 | 5507 | 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 | 5508 | 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 | 5509 | 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 | 5510 | 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 | 5511 | 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 | 5512 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5513 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5514 | 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 | 5515 | 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 | 5516 |   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 | 5517 | 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 | 5518 | 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 | 5519 | 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 | 5520 | 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 | 5521 | 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 | 5522 | 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 | 5523 | 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 | 5524 | 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 | 5525 | 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 | 5526 | 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 | 5527 | 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 | 5528 | 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 | 5529 | 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 | 5530 | 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 | 5531 | 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 | 5532 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5533 | apply(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 | 5534 | 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 | 5535 | 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 | 5536 | 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 | 5537 | 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 | 5538 | 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 | 5539 | 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 | 5540 | 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 | 5541 | 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 | 5542 | 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 | 5543 | 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 | 5544 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 5546 | 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 | 5547 | 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 | 5548 | 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 | 5549 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5550 | apply(simp 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 | 5551 | 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 | 5552 | 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 | 5553 | 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 | 5554 | 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 | 5555 | 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 | 5556 | 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 | 5557 | 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 | 5558 | 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 | 5559 | 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 | 5560 | 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 | 5561 | 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 | 5562 | 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 | 5563 | 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 | 5564 | 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 | 5565 | 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 | 5566 | 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 | 5567 | 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 | 5568 | 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 | 5569 | 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 | 5570 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5571 | apply(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 | 5572 | 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 | 5573 | 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 | 5574 | 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 | 5575 | 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 | 5576 | 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 | 5577 | 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 | 5578 | 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 | 5579 | 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 | 5580 | 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 | 5581 | 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 | 5582 | 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 | 5583 | 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 | 5584 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5585 | apply(simp 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 | 5586 | 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 | 5587 | 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 | 5588 | 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 | 5589 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5592 |   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 | 5593 | 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 | 5594 | 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 | 5595 | 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 | 5596 | 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 | 5597 | 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 | 5598 | 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 | 5599 | 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 | 5600 | 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 | 5601 | 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 | 5602 | 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 | 5603 | 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 | 5604 | 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 | 5605 | 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 | 5606 | 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 | 5607 | 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 | 5608 | 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 | 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_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 | 5613 | 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 | 5614 | 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 | 5615 | 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 | 5616 | 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 | 5617 | 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 | 5618 | 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 | 5619 | 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 | 5620 | 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 | 5621 | 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 | 5622 | 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 | 5623 | 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 | 5624 | 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 | 5625 | 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 | 5626 | 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 | 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 5632 | 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 | 5633 | 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 | 5634 | 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 | 5635 | 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 | 5636 | 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 | 5637 | 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 | 5638 | 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 | 5639 | 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 | 5640 | 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 | 5641 | 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 | 5642 | 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 | 5643 | 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 | 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 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 | 5647 | 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 | 5648 |   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 | 5649 | 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 | 5650 | 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 | 5651 | 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 | 5652 | 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 | 5653 | 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 | 5654 | 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 | 5655 | 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 | 5656 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 5658 | 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 | 5659 | 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 | 5660 | 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 | 5661 | 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 | 5662 | 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 | 5663 | 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 | 5664 | 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 | 5665 | 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 | 5666 | 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 | 5667 | 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 | 5668 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 5670 | 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 | 5671 | 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 | 5672 | 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 | 5673 | 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 | 5674 | 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 | 5675 | 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 | 5676 | 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 | 5677 | 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 | 5678 | 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 | 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(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 | 5681 | 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 | 5682 | 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 | 5683 | 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 | 5684 | 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 | 5685 | 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 | 5686 | 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 | 5687 | 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 | 5688 | 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 | 5689 | 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 | 5690 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 5692 | 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 | 5693 | 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 | 5694 | 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 | 5695 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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: 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(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 | 5698 | 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 | 5699 | 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 | 5700 | 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 | 5701 | 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 | 5702 | 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 | 5703 | 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 | 5704 | 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 | 5705 | 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 | 5706 | 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 | 5707 | 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 | 5708 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5709 | apply(simp 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 | 5710 | 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 | 5711 | 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 | 5712 | 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 | 5713 | 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 | 5714 | 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 | 5715 | 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 | 5716 | 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 | 5717 | 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 | 5718 | 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 | 5719 | 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 | 5720 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* Reductions *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 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 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 | 5725 | assumes a: "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 | 5726 | and b: "M \<longrightarrow>\<^isub>l 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 | 5727 | 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 | 5728 | 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 | 5729 | 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 | 5730 | 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 | 5731 | 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 | 5732 | 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 | 5733 | 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 | 5734 | 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 | 5735 | 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 | 5736 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5739 | assumes a: "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 | 5740 | and b: "M \<longrightarrow>\<^isub>c 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 | 5741 | 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 | 5742 | 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 | 5743 | 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 | 5744 | 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 | 5745 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5748 | assumes a: "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 | 5749 | and b: "M \<longrightarrow>\<^isub>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 | 5750 | 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 | 5751 | 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 | 5752 | 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 | 5753 | 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 | 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(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 | 5756 | 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 | 5757 | 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 | 5758 | 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 | 5759 | 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 | 5760 | 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 | 5761 | 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 | 5762 | 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 | 5763 | 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 | 5764 | 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 | 5765 | 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 | 5766 | 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 | 5767 | 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 | 5768 | 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 | 5769 | 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 | 5770 | 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 | 5771 | 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 | 5772 | 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 | 5773 | 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 | 5774 | 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 | 5775 | 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 | 5776 | 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 | 5777 | 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 | 5778 | 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 | 5779 | 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 | 5780 | 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 | 5781 | 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 | 5782 | 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 | 5783 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5786 | assumes a: "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 | 5787 | and b: "M \<longrightarrow>\<^isub>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 | 5788 | 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 | 5789 | 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 | 5790 | 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 | 5791 | apply(auto simp add: 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 | 5792 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5795 | assumes a: "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 | 5796 | and b: "M \<longrightarrow>\<^isub>l 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 | 5797 | 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 | 5798 | 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 | 5799 | 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 | 5800 | 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 | 5801 | 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 | 5802 | 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 | 5803 | 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 | 5804 | 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 | 5805 | 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 | 5806 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5809 | assumes a: "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 | 5810 | and b: "M \<longrightarrow>\<^isub>c 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 | 5811 | 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 | 5812 | 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 | 5813 | 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 | 5814 | 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 | 5815 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5818 | assumes a: "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 | 5819 | and b: "M \<longrightarrow>\<^isub>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 | 5820 | 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 | 5821 | 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 | 5822 | 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 | 5823 | 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 | 5824 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 5826 | 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 | 5827 | 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 | 5828 | 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 | 5829 | 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 | 5830 | 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 | 5831 | 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 | 5832 | 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 | 5833 | 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 | 5834 | 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 | 5835 | 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 | 5836 | 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 | 5837 | 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 | 5838 | 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 | 5839 | 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 | 5840 | 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 | 5841 | 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 | 5842 | 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 | 5843 | 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 | 5844 | 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 | 5845 | 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 | 5846 | 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 | 5847 | 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 | 5848 | 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 | 5849 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 5852 | assumes a: "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 | 5853 | and b: "M \<longrightarrow>\<^isub>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 | 5854 | 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 | 5855 | 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 | 5856 | 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 | 5857 | apply(auto simp add: 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 | 5858 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5859 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | text {* substitution properties *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma subst_with_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 | 5863 |   shows "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* 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 | 5864 | 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 | 5865 | case (Ax z 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 | 5866 |   show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[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 | 5867 | 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 | 5868 | 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 | 5869 | 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 | 5870 |     have "(Ax z b){x:=<a>.Ax y a} = Cut <a>.Ax y a (x).Ax x 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 | 5871 | also have "\<dots> \<longrightarrow>\<^isub>a* (Ax x b)[x\<turnstile>n>y]" 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 | 5872 |     finally show "Ax z b{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]" 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 | 5873 | 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 | 5874 | 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 | 5875 | 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 | 5876 |     then show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]" using 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 | 5877 | 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 | 5878 | 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 | 5879 | 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 | 5880 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 5882 |   have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" 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 | 5883 |   show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]"
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5884 | 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 | 5885 | 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 | 5886 | 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 | 5887 |     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
 | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5888 | also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.Ax y a (z).(N[x\<turnstile>n>y])" using ih2 a_star_congs 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 | 5889 | 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 | 5890 | by (simp add: trm.inject 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 | 5891 |     finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" 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 | 5892 | 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 | 5893 | 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 | 5894 | 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 | 5895 |     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 | 5896 | 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 | 5897 | also have "\<dots> \<longrightarrow>\<^isub>a* Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using ih1 ih2 a_star_congs 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 | 5898 |     finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" 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 | 5899 | 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 | 5900 | 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 | 5901 | 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 | 5902 | have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>b" 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 | 5903 |   have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 5904 |   have "(NotR (z).M b){x:=<a>.Ax y a} = NotR (z).(M{x:=<a>.Ax y a}) b" 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 | 5905 | also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[x\<turnstile>n>y]) b" using ih 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 | 5906 |   finally show "(NotR (z).M b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotR (z).M b)[x\<turnstile>n>y]" 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 | 5907 | 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 | 5908 | 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 | 5909 | have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>z" 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 | 5910 |   have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 5911 |   show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[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 | 5912 | 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 | 5913 | 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 | 5914 | 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 | 5915 |     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 | 5916 |     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 | 5917 |                         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 | 5918 | 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 | 5919 |     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 | 5920 | using new by (simp add: fresh_fun_simp_NotL fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5921 |     also have "\<dots> \<longrightarrow>\<^isub>a* (NotL <b>.(M{x:=<a>.Ax y a}) x')[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 | 5922 | 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 | 5923 | 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 | 5924 | 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 | 5925 | 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 | 5926 | 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 | 5927 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = NotL <b>.(M{x:=<a>.Ax y a}) y" using new by (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 | 5929 | also have "\<dots> \<longrightarrow>\<^isub>a* NotL <b>.(M[x\<turnstile>n>y]) y" using ih 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 | 5930 | also have "\<dots> = (NotL <b>.M z)[x\<turnstile>n>y]" 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 | 5931 |     finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]" 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 | 5932 | 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 | 5933 | 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 | 5934 | 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 | 5935 |     have "(NotL <b>.M z){x:=<a>.Ax y a} = NotL <b>.(M{x:=<a>.Ax y a}) z" 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 | 5936 | also have "\<dots> \<longrightarrow>\<^isub>a* NotL <b>.(M[x\<turnstile>n>y]) z" using ih 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 | 5937 |     finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]" using 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 | 5938 | 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 | 5939 | 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 | 5940 | 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 | 5941 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 5943 |   have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" 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 | 5944 |   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 | 5945 | 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 | 5946 | also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[x\<turnstile>n>y]) <d>.(N[x\<turnstile>n>y]) e" 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 | 5947 |   finally show "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[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 | 5948 | 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 | 5949 | 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 | 5950 | 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 | 5951 | have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" 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 | 5952 |   have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 5953 |   show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[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 | 5954 | 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 | 5955 | 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 | 5956 | 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 | 5957 |     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 | 5958 |     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 | 5959 |                         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 | 5960 | 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 | 5961 |     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 | 5962 | using new by (simp add: fresh_fun_simp_AndL1 fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 5963 |     also have "\<dots> \<longrightarrow>\<^isub>a* (AndL1 (u).(M{x:=<a>.Ax y a}) v')[v'\<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 | 5964 | 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 | 5965 | 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 | 5966 | 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 | 5967 | 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 | 5968 | 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 | 5969 | 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 | 5970 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = AndL1 (u).(M{x:=<a>.Ax y a}) y" using fs 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 | 5972 | by (auto simp add: fresh_prod fresh_atm 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 | 5973 | also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[x\<turnstile>n>y]) y" using ih 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 | 5974 | also have "\<dots> = (AndL1 (u).M v)[x\<turnstile>n>y]" 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 | 5975 |     finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" 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 | 5976 | 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 | 5977 | 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 | 5978 | 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 | 5979 |     have "(AndL1 (u).M v){x:=<a>.Ax y a} = AndL1 (u).(M{x:=<a>.Ax y a}) v" 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 | 5980 | also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[x\<turnstile>n>y]) v" using ih 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 | 5981 |     finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" 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 | 5982 | 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 | 5983 | 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 | 5984 | 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 | 5985 | have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" 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 | 5986 |   have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 5987 |   show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[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 | 5988 | 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 | 5989 | 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 | 5990 | 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 | 5991 |     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 | 5992 |     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 | 5993 |                         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 | 5994 | 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 | 5995 |     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 | 5996 | using new by (simp add: fresh_fun_simp_AndL2 fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> \<longrightarrow>\<^isub>a* (AndL2 (u).(M{x:=<a>.Ax y a}) v')[v'\<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 | 5998 | 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 | 5999 | 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 | 6000 | 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 | 6001 | 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 | 6002 | 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 | 6003 | 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 | 6004 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = AndL2 (u).(M{x:=<a>.Ax y a}) y" using fs 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 | 6006 | by (auto simp add: fresh_prod fresh_atm 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 | 6007 | also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[x\<turnstile>n>y]) y" using ih 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 | 6008 | also have "\<dots> = (AndL2 (u).M v)[x\<turnstile>n>y]" 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 | 6009 |     finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" 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 | 6010 | 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 | 6011 | 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 | 6012 | 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 | 6013 |     have "(AndL2 (u).M v){x:=<a>.Ax y a} = AndL2 (u).(M{x:=<a>.Ax y a}) v" 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 | 6014 | also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[x\<turnstile>n>y]) v" using ih 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 | 6015 |     finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" 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 | 6016 | 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 | 6017 | 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 | 6018 | 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 | 6019 | have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" 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 | 6020 |   have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 6021 |   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)
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[x\<turnstile>n>y]) d" using ih 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 | 6023 |   finally show "(OrR1 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[x\<turnstile>n>y]" 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 | 6024 | 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 | 6025 | 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 | 6026 | have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" 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 | 6027 |   have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 6028 |   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)
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[x\<turnstile>n>y]) d" using ih 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 | 6030 |   finally show "(OrR2 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[x\<turnstile>n>y]" 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 | 6031 | 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 | 6032 | 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 | 6033 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 6035 |   have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" 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 | 6036 |   show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[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 | 6037 | 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 | 6038 | 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 | 6039 | 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 | 6040 |     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 | 6041 | 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 | 6042 |     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 | 6043 |                  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 | 6044 | 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 | 6045 |     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 | 6046 | using new by (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 | 6047 |     also have "\<dots> \<longrightarrow>\<^isub>a* (OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z')[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 | 6048 | 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 | 6049 | 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 | 6050 | 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 | 6051 | 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 | 6052 | 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 | 6053 | 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 | 6054 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6055 |     also have "\<dots> = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) y" using fs 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 | 6056 | by (auto simp add: fresh_prod fresh_atm nrename_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 | 6057 | also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>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 | 6058 | 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 | 6059 | also have "\<dots> = (OrL (u).M (v).N z)[x\<turnstile>n>y]" 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 | 6060 |     finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" 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 | 6061 | 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 | 6062 | 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 | 6063 | 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 | 6064 |     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 | 6065 | 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 | 6066 | also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) z" | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6067 | 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 | 6068 |     finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" 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 | 6069 | 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 | 6070 | 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 | 6071 | 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 | 6072 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 6074 |   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
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> \<longrightarrow>\<^isub>a* ImpR (z).<c>.(M[x\<turnstile>n>y]) d" using ih 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 | 6076 |   finally show "(ImpR (z).<c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[x\<turnstile>n>y]" 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 | 6077 | 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 | 6078 | 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 | 6079 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" 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 | 6081 |   have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" 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 | 6082 |   show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[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 | 6083 | 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 | 6084 | 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 | 6085 | 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 | 6086 |     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 | 6087 | 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 | 6088 |     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 | 6089 |                  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 | 6090 | 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 | 6091 |     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 | 6092 | using new by (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 | 6093 |     also have "\<dots> \<longrightarrow>\<^isub>a* (ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v')[v'\<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 | 6094 | 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 | 6095 | 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 | 6096 | 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 | 6097 | 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 | 6098 | 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 | 6099 | 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 | 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 |     also have "\<dots> = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) y" using fs 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 | 6102 | by (auto simp add: fresh_prod subst_fresh fresh_atm trm.inject alpha 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 | 6103 | also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>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 | 6104 | 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 | 6105 | also have "\<dots> = (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" 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 | 6106 |     finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" 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 | 6107 | 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 | 6108 | 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 | 6109 | 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 | 6110 |     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 | 6111 | 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 | 6112 | also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) 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 | 6113 | 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 | 6114 |     finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[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 | 6115 | 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 | 6116 | 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 | 6117 | 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 | 6118 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | lemma subst_with_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 | 6120 |   shows "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* 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 | 6121 | 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 | 6122 | case (Ax z 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 | 6123 |   show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[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 | 6124 | 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 | 6125 | 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 | 6126 | 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 | 6127 |     have "(Ax z c){b:=(x).Ax x a} = Cut <b>.Ax z c (x).Ax x a" 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 | 6128 | also have "\<dots> \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" using eq 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 | 6129 |     finally show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" 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 | 6130 | 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 | 6131 | 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 | 6132 | 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 | 6133 |     then show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" 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 | 6134 | 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 | 6135 | 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 | 6136 | 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 | 6137 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6139 |   have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" 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 | 6140 |   show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[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 | 6141 | 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 | 6142 | 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 | 6143 | 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 | 6144 |     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 
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> \<longrightarrow>\<^isub>a* Cut <c>.(M[b\<turnstile>c>a]) (x).Ax x a" using ih1 a_star_congs 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 | 6146 | 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 | 6147 | by (simp add: trm.inject 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 | 6148 |     finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" 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 | 6149 | 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 | 6150 | 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 | 6151 | 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 | 6152 |     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 | 6153 | 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 | 6154 | also have "\<dots> \<longrightarrow>\<^isub>a* Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using ih1 ih2 a_star_congs 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 | 6155 |     finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" 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 | 6156 | 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 | 6157 | 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 | 6158 | 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 | 6159 | have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>c" 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 | 6160 |   have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6161 |   show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[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 | 6162 | 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 | 6163 | 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 | 6164 | 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 | 6165 |     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 | 6166 |     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 | 6167 |                         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 | 6168 | 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 | 6169 |     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 | 6170 | using new by (simp add: fresh_fun_simp_NotR fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> \<longrightarrow>\<^isub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\<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 | 6172 | 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 | 6173 | 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 | 6174 | 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 | 6175 | 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 | 6176 | 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 | 6177 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6179 |     also have "\<dots> = NotR (z).(M{b:=(x).Ax x a}) a" using new by (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 | 6180 | also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[b\<turnstile>c>a]) a" using ih 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 | 6181 | also have "\<dots> = (NotR (z).M c)[b\<turnstile>c>a]" 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 | 6182 |     finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]" 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 | 6183 | 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 | 6184 | 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 | 6185 | 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 | 6186 |     have "(NotR (z).M c){b:=(x).Ax x a} = NotR (z).(M{b:=(x).Ax x a}) c" 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 | 6187 | also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[b\<turnstile>c>a]) c" using ih 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 | 6188 |     finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]" using 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 | 6189 | 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 | 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 (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 | 6192 | have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>z" 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 | 6193 |   have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6194 |   have "(NotL <c>.M z){b:=(x).Ax x a} = NotL <c>.(M{b:=(x).Ax x a}) z" 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 | 6195 | also have "\<dots> \<longrightarrow>\<^isub>a* NotL <c>.(M[b\<turnstile>c>a]) z" using ih 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 | 6196 |   finally show "(NotL <c>.M z){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotL <c>.M z)[b\<turnstile>c>a]" 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 | 6197 | 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 | 6198 | 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 | 6199 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6201 |   have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" 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 | 6202 |   show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[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 | 6203 | 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 | 6204 | 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 | 6205 | 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 | 6206 |     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 | 6207 | 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 | 6208 |     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 | 6209 |                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 | 6210 | 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 | 6211 |     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 | 6212 | using new by (simp add: fresh_fun_simp_AndR fresh_prod) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> \<longrightarrow>\<^isub>a* (AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e')[e'\<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 | 6214 | 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 | 6215 | 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 | 6216 | 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 | 6217 | 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 | 6218 | 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 | 6219 | 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 | 6220 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) a" using fs 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 | 6222 | by (auto simp add: fresh_prod fresh_atm subst_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 | 6223 | also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) a" 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 | 6224 | also have "\<dots> = (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" 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 | 6225 |     finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" 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 | 6226 | 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 | 6227 | 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 | 6228 | 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 | 6229 |     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 | 6230 | 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 | 6231 | also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) e" 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 | 6232 |     finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[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 | 6233 | 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 | 6234 | 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 | 6235 | 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 | 6236 | 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 | 6237 | have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" 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 | 6238 |   have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6239 |   have "(AndL1 (u).M v){b:=(x).Ax x a} = AndL1 (u).(M{b:=(x).Ax x a}) v" 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 | 6240 | also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[b\<turnstile>c>a]) v" using ih 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 | 6241 |   finally show "(AndL1 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[b\<turnstile>c>a]" 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 | 6242 | 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 | 6243 | 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 | 6244 | have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" 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 | 6245 |   have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6246 |   have "(AndL2 (u).M v){b:=(x).Ax x a} = AndL2 (u).(M{b:=(x).Ax x a}) v" 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 | 6247 | also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[b\<turnstile>c>a]) v" using ih 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 | 6248 |   finally show "(AndL2 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[b\<turnstile>c>a]" 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 | 6249 | 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 | 6250 | 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 | 6251 | have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" 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 | 6252 |   have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6253 |   show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[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 | 6254 | 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 | 6255 | 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 | 6256 | 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 | 6257 |     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 | 6258 | 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 | 6259 |     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 | 6260 |              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 | 6261 |     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 | 6262 | using new by (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 | 6263 |     also have "\<dots> \<longrightarrow>\<^isub>a* (OrR1 <c>.M{b:=(x).Ax x a} a')[a'\<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 | 6264 | 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 | 6265 | 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 | 6266 | 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 | 6267 | 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 | 6268 | 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 | 6269 | 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 | 6270 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = OrR1 <c>.M{b:=(x).Ax x a} a" using fs 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 | 6272 | by (auto simp add: fresh_prod fresh_atm crename_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 | 6273 | also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[b\<turnstile>c>a]) a" using ih 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 | 6274 | also have "\<dots> = (OrR1 <c>.M d)[b\<turnstile>c>a]" 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 | 6275 |     finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" 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 | 6276 | 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 | 6277 | 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 | 6278 | 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 | 6279 |     have "(OrR1 <c>.M d){b:=(x).Ax x a} = OrR1 <c>.(M{b:=(x).Ax x a}) d" 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 | 6280 | also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[b\<turnstile>c>a]) d" using ih 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 | 6281 |     finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" 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 | 6282 | 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 | 6283 | 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 | 6284 | 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 | 6285 | have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" 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 | 6286 |   have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6287 |   show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[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 | 6288 | 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 | 6289 | 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 | 6290 | 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 | 6291 |     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 | 6292 | 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 | 6293 |     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 | 6294 |              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 | 6295 |     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 | 6296 | using new by (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 | 6297 |     also have "\<dots> \<longrightarrow>\<^isub>a* (OrR2 <c>.M{b:=(x).Ax x a} a')[a'\<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 | 6298 | 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 | 6299 | 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 | 6300 | 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 | 6301 | 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 | 6302 | 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 | 6303 | 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 | 6304 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = OrR2 <c>.M{b:=(x).Ax x a} a" using fs 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 | 6306 | by (auto simp add: fresh_prod fresh_atm crename_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 | 6307 | also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[b\<turnstile>c>a]) a" using ih 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 | 6308 | also have "\<dots> = (OrR2 <c>.M d)[b\<turnstile>c>a]" 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 | 6309 |     finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" 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 | 6310 | 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 | 6311 | 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 | 6312 | 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 | 6313 |     have "(OrR2 <c>.M d){b:=(x).Ax x a} = OrR2 <c>.(M{b:=(x).Ax x a}) d" 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 | 6314 | also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[b\<turnstile>c>a]) d" using ih 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 | 6315 |     finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" 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 | 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 | 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 | 6318 | 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 | 6319 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6321 |   have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" 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 | 6322 |   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 | 6323 | 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 | 6324 | also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[b\<turnstile>c>a]) (v).(N[b\<turnstile>c>a]) z" 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 | 6325 |   finally show "(OrL (u).M (v).N z){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[b\<turnstile>c>a]" 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 | 6326 | 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 | 6327 | 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 | 6328 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6330 |   show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[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 | 6331 | 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 | 6332 | 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 | 6333 | 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 | 6334 |     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 | 6335 | 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 | 6336 |     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 | 6337 |                 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 | 6338 |     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 | 6339 | using new by (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 | 6340 |     also have "\<dots> \<longrightarrow>\<^isub>a* (ImpR z.<c>.M{b:=(x).Ax x a} a')[a'\<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 | 6341 | 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 | 6342 | 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 | 6343 | 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 | 6344 | 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 | 6345 | 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 | 6346 | 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 | 6347 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = ImpR z.<c>.M{b:=(x).Ax x a} a" using fs 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 | 6349 | by (auto simp add: fresh_prod crename_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 | 6350 | also have "\<dots> \<longrightarrow>\<^isub>a* ImpR z.<c>.(M[b\<turnstile>c>a]) a" using ih 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 | 6351 | also have "\<dots> = (ImpR z.<c>.M b)[b\<turnstile>c>a]" 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 | 6352 |     finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" 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 | 6353 | 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 | 6354 | 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 | 6355 | 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 | 6356 |     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
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | also have "\<dots> \<longrightarrow>\<^isub>a* ImpR (z).<c>.(M[b\<turnstile>c>a]) d" using ih 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 | 6358 |     finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using neq 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 | 6359 | 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 | 6360 | 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 | 6361 | 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 | 6362 | 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+ | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |   have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" 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 | 6364 |   have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" 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 | 6365 |   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 | 6366 | 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 | 6367 | also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[b\<turnstile>c>a]) (u).(N[b\<turnstile>c>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 | 6368 | 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 | 6369 |   finally show "(ImpL <c>.M (u).N v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[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 | 6370 | 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 | 6371 | 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 | 6372 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6373 | text {* substitution lemmas *}
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6376 |   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 | 6377 | apply(nominal_induct M avoiding: b y Q x a rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6378 | apply(auto simp add: fresh_atm 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 | 6379 | 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 | 6380 | 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 | 6381 | 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 | 6382 | 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 | 6383 | 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 | 6384 | 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 | 6385 | 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 | 6386 | 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 | 6387 | 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 | 6388 | 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 | 6389 | 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 | 6390 | 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 | 6391 | 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 | 6392 | 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 | 6393 | 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 | 6394 | 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 | 6395 | 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 | 6396 | 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 | 6397 | 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 | 6398 | 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 | 6399 | 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 | 6400 | 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 | 6401 | 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 | 6402 | 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 | 6403 | 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 | 6404 | 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 | 6405 | 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 | 6406 | 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 | 6407 | 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 | 6408 | 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 | 6409 | 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 | 6410 | 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 | 6411 | 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 | 6412 | 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 | 6413 | 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 | 6414 | 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 | 6415 | 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 | 6416 | 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 | 6417 | 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 | 6418 | 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 | 6419 | 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 | 6420 | 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 | 6421 | 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 | 6422 | 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 | 6423 | 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 | 6424 | 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 | 6425 | 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 | 6426 | 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 | 6427 | 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 | 6428 | 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 | 6429 | 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 | 6430 | 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 | 6431 | 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 | 6432 | 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 | 6433 | 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 | 6434 | 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 | 6435 | 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 | 6436 | 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 | 6437 | 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 | 6438 | 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 | 6439 | 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 | 6440 | 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 | 6441 | 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 | 6442 | 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 | 6443 | 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 | 6444 | 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 | 6445 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6448 |   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 | 6449 | apply(nominal_induct M avoiding: b y Q x a rule: trm.strong_induct) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6450 | apply(auto simp add: fresh_atm 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 | 6451 | 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 | 6452 | 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 | 6453 | 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 | 6454 | 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 | 6455 | 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 | 6456 | 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 | 6457 | 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 | 6458 | 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 | 6459 | 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 | 6460 | 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 | 6461 | 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 | 6462 | 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 | 6463 | 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 | 6464 | 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 | 6465 | 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 | 6466 | 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 | 6467 | 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 | 6468 | 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 | 6469 | 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 | 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_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 | 6474 | 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 | 6475 | 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 | 6476 | 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 | 6477 | 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 | 6478 | 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 | 6479 | 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 | 6480 | 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 | 6481 | 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 | 6482 | 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 | 6483 | 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 | 6484 | 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 | 6485 | 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 | 6486 | 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 | 6487 | 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 | 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_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 | 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 | 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 | 6494 | 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 | 6495 | 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 | 6496 | 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 | 6497 | 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 | 6498 | 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 | 6499 | 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 | 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_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 | 6504 | 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 | 6505 | 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 | 6506 | 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 | 6507 | 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 | 6508 | 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 | 6509 | 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 | 6510 | 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 | 6511 | 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 | 6512 | 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 | 6513 | 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 | 6514 | 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 | 6515 | 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 | 6516 | 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 | 6517 | 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 | 6518 | 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 | 6519 | 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 | 6520 | 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 | 6521 | 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 | 6522 | 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 | 6523 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6525 | 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 | 6526 | 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 | 6527 |   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 | 6528 | 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 | 6529 | 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 | 6530 | 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 | 6531 | 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 | 6532 | by (auto simp add: abs_fresh fresh_atm forget 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 | 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 (Cut d M u M' x' y' c P) | 
| 41893 | 6535 | 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 | 6536 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6538 | 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 | 6539 | 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 | 6540 | 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 | 6541 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6544 | 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 | 6545 | 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 | 6546 | 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 | 6547 | 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 | 6548 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6551 | 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 | 6552 | 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 | 6553 | 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 | 6554 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6557 | 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 | 6558 | 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 | 6559 | 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 | 6560 | 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 | 6561 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6564 | 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 | 6565 | 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 | 6566 | 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 | 6567 | 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 | 6568 | 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 | 6569 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6572 | 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 | 6573 | 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 | 6574 | 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 | 6575 | 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 | 6576 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6579 | 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 | 6580 | 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 | 6581 | 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 | 6582 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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: 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 | 6585 | 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 | 6586 | 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 | 6587 | 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 | 6588 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6591 | 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 | 6592 | 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 | 6593 | 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 | 6594 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6597 | 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 | 6598 | 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 | 6599 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6602 | 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 | 6603 | 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 | 6604 | 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 | 6605 | 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 | 6606 | 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 | 6607 | 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 | 6608 | 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 | 6609 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6612 | 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 | 6613 | 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 | 6614 | 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 | 6615 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 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 | by (auto simp add: abs_fresh fresh_atm 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 | 6620 | 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 | 6621 | 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 | 6622 | 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 | 6623 | apply (auto simp add: abs_fresh fresh_atm 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 | 6624 |     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 | 6625 | 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 | 6626 | 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 | 6627 | 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 | 6628 | 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 | 6629 | 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 | 6630 | 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 | 6631 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 6634 | 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 | 6635 | 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 | 6636 |     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 | 6637 | 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 | 6638 | 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 | 6639 | 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 | 6640 | 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 | 6641 | 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 | 6642 | 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 | 6643 | 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 | 6644 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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: 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 | 6647 | 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 | 6648 | 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 | 6649 | 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 | 6650 | 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 | 6651 | 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 | 6652 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6654 | 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 | 6655 | 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 | 6656 | by (auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6657 | 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 | 6658 | 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 | 6659 | 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 | 6660 | apply(auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6661 |     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 | 6662 | 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 | 6663 | 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 | 6664 | 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 | 6665 | 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 | 6666 | 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 | 6667 | 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 | 6668 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(auto 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 | 6671 | 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 | 6672 | 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 | 6673 |     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 | 6674 | 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 | 6675 | 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 | 6676 | 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 | 6677 | 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 | 6678 | 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 | 6679 | 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 | 6680 | 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 | 6681 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 6684 | 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 | 6685 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6687 | 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 | 6688 | 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 | 6689 | apply(auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6690 |     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 | 6691 | 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 | 6692 | 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 | 6693 | 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 | 6694 | 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 | 6695 | 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 | 6696 | 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 | 6697 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 6700 | 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 | 6701 | 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 | 6702 |     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 | 6703 | 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 | 6704 | 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 | 6705 | 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 | 6706 | 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 | 6707 | 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 | 6708 | 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 | 6709 | 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 | 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(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(auto 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 | 6713 | 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 | 6714 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6716 | 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 | 6717 | 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 | 6718 | by (auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6719 | 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 | 6720 | 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 | 6721 | 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 | 6722 | by (auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6723 | 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 | 6724 | 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 | 6725 | 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 | 6726 | apply(auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6727 |     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 | 6728 |                                         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 | 6729 | 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 | 6730 | 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 | 6731 | 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 | 6732 | 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 | 6733 | 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 | 6734 | 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 | 6735 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(auto 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 | 6738 | 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 | 6739 | 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 | 6740 | 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 | 6741 | 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 | 6742 | 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 | 6743 | 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 | 6744 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 6746 |     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 | 6747 |                                         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 | 6748 | 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 | 6749 | 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 | 6750 | 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 | 6751 | 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 | 6752 | 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 | 6753 | 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 | 6754 | 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 | 6755 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6758 | 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 | 6759 | 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 | 6760 | 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 | 6761 | 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 | 6762 | 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 | 6763 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 6765 | 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 | 6766 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6768 | 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 | 6769 | 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 | 6770 | by (auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6771 | 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 | 6772 | 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 | 6773 | 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 | 6774 | apply(auto simp add: abs_fresh fresh_atm forget trm.inject 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(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 | 6776 |                                         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 | 6777 | 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 | 6778 | 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 | 6779 | 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 | 6780 | 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 | 6781 | 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 | 6782 | 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 | 6783 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 6786 | 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 | 6787 | 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 | 6788 | 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 | 6789 | 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 | 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(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 | 6792 | 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 | 6793 |     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: 
41893diff
changeset | 6794 |                                         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 | 6795 | 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 | 6796 | 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 | 6797 | 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 | 6798 | 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 | 6799 | 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 | 6800 | 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 | 6801 | 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 | 6802 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6805 | 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 | 6806 | 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 | 6807 | 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 | 6808 | 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 | 6809 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 6811 | 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 | 6812 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6814 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6816 | 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 | 6817 |   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 | 6818 | 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 | 6819 | 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 | 6820 | 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 | 6821 | 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 | 6822 | 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 | 6823 | 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 | 6824 | 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 | 6825 |       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 | 6826 | 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 | 6827 | 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 | 6828 | 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 | 6829 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6831 | 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 | 6832 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6834 | 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 | 6835 |   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 | 6836 | 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 | 6837 | 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 | 6838 | 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 | 6839 | 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 | 6840 | by (auto simp add: abs_fresh fresh_atm forget 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 | 6841 | 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 | 6842 | case (Cut d M u M' x' y' c P) | 
| 41893 | 6843 | 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 | 6844 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(auto 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 | 6846 | 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 | 6847 | 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 | 6848 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6850 | 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 | 6851 | 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 | 6852 | 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 | 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_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 | 6855 | 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 | 6856 | 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 | 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(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 | 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 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 | 6861 | 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 | 6862 | 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 | 6863 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 6865 | 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 | 6866 | 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 | 6867 | 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 | 6868 | 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 | 6869 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6870 | 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 | 6871 | 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 | 6872 | 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 | 6873 | 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 | 6874 | 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 | 6875 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 6877 | 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 | 6878 | 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 | 6879 | 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 | 6880 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6882 | 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 | 6883 | 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 | 6884 | 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 | 6885 | 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 | 6886 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6888 | 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 | 6889 | 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 | 6890 | 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 | 6891 | 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 | 6892 | 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 | 6893 | 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 | 6894 | 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 | 6895 | 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 | 6896 | 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 | 6897 | 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 | 6898 | 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 | 6899 | 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 | 6900 | 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 | 6901 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 6903 | 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 | 6904 | 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 | 6905 | by (auto simp add: abs_fresh fresh_atm 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 | 6906 | 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 | 6907 | 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 | 6908 | 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 | 6909 | apply (auto simp add: abs_fresh fresh_atm 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 | 6910 |     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 | 6911 | 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 | 6912 | 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 | 6913 | 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 | 6914 | 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 | 6915 | 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 | 6916 | 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 | 6917 | 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 | 6918 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 6920 | 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 | 6921 | 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 | 6922 |     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 | 6923 | 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 | 6924 | 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 | 6925 | 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 | 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_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 | 6929 | 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 | 6930 | 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 | 6931 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6933 | 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 | 6934 | 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 | 6935 | 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 | 6936 | 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 | 6937 | 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 | 6938 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 6939 | 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 | 6940 | 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 | 6941 | 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 | 6942 | apply(auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6943 |     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 | 6944 |                                         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 | 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_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 | 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_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 | 6950 | 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 | 6951 | 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 | 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(auto 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 | 6954 | 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 | 6955 | 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 | 6956 | 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 | 6957 | 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 | 6958 | 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 | 6959 | 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 | 6960 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 6962 |     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 | 6963 |                                         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 | 6964 | 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 | 6965 | 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 | 6966 | 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 | 6967 | 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 | 6968 | 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 | 6969 | 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 | 6970 | 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 | 6971 | 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 | 6972 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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: 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 | 6974 | 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 | 6975 | 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 | 6976 | 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 | 6977 | 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 | 6978 | 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 | 6979 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | apply(auto 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 | 6981 | 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 | 6982 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 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 | 6985 | 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 | 6986 | by (auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6987 | 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 | 6988 | 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 | 6989 | 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 | 6990 | by (auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 6991 | 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 | 6992 | 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 | 6993 | 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 | 6994 | apply (auto simp add: abs_fresh fresh_atm 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 | 6995 |     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 | 6996 | 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 | 6997 | 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 | 6998 | 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 | 6999 | 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 | 7000 | 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 | 7001 | 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 | 7002 | 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 | 7003 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 7005 | 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 | 7006 | 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 | 7007 |     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 | 7008 | 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 | 7009 | 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 | 7010 | 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 | 7011 | 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 | 7012 | 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 | 7013 | 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 | 7014 | 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 | 7015 | 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 | 7016 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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: 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 | 7018 | 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 | 7019 | 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 | 7020 | 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 | 7021 | 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 | 7022 | 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 | 7023 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7025 | 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 | 7026 | 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 | 7027 | apply (auto simp add: abs_fresh fresh_atm 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 | 7028 |     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 | 7029 | 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 | 7030 | 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 | 7031 | 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 | 7032 | 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 | 7033 | 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 | 7034 | 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 | 7035 | 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 | 7036 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 7038 | 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 | 7039 | 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 | 7040 |     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 | 7041 | 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 | 7042 | 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 | 7043 | 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 | 7044 | 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 | 7045 | 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 | 7046 | 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 | 7047 | 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 | 7048 | 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 | 7049 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 7051 | 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 | 7052 | 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 | 7053 | 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 | 7054 | 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 | 7055 | 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 | 7056 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7058 | 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 | 7059 | 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 | 7060 | by(auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 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 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 | 7063 | 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 | 7064 | by (auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 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 (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 | 7067 | 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 | 7068 | apply(auto simp add: abs_fresh fresh_atm forget trm.inject 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 | 7069 |     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 | 7070 | 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 | 7071 | 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 | 7072 | 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 | 7073 | 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 | 7074 | 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 | 7075 | 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 | 7076 | 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 | 7077 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 7079 | 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 | 7080 | 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 | 7081 |     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 | 7082 | 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 | 7083 | 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 | 7084 | 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 | 7085 | 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 | 7086 | 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 | 7087 | 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 | 7088 | 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 | 7089 | 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 | 7090 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(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 | 7092 | 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 | 7093 | 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 | 7094 | 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 | 7095 | 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 | 7096 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to 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(auto 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 | 7098 | 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 | 7099 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7101 | |
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7103 | 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 | 7104 |   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 | 7105 | 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 | 7106 | 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 | 7107 | 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 | 7108 | 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 | 7109 | 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 | 7110 | 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 | 7111 | 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 | 7112 |       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 | 7113 | 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 | 7114 | 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 | 7115 | 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 | 7116 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7118 | 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 | 7119 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7120 | 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 | 7121 | 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 | 7122 |   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 | 7123 | 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 | 7124 | 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 | 7125 | 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 | 7126 | 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 | 7127 |   { 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 | 7128 |     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 | 7129 |     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 | 7130 | 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 | 7131 |     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 | 7132 |     also have "\<dots> = (Cut <b>.Ax x 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 | 7133 | using fs asm by (auto 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 | 7134 |     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 | 7135 |     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 | 7136 | 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 | 7137 | } | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7139 |   { 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 | 7140 |     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 | 7141 | 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 | 7142 |     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 | 7143 | 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 | 7144 |     also have "\<dots> = (Cut <b>.Ax z c (y).Q){x:=<a>.(P{b:=(y).Q})}" using asm 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 | 7145 | by (auto simp add: trm.inject subst_fresh 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 | 7146 |     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 | 7147 |     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 | 7148 | } | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7150 |   { 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 | 7151 |     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
 | 
| 
9be4ab2acc13
split Class.thy into parts to 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 |     also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (x).Ax z c" using fs asm by (auto simp add: 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 | 7153 |     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 | 7154 |     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 | 7155 |     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 | 7156 | } | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7158 |   { 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 | 7159 |     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 | 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 | 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 | 7162 | 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 | 7163 | 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 | 7164 |   { 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 | 7165 |     have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (z).(N{x:=<a>.P})){b:=(y).Q}" 
 | 
| 41893 | 7166 | using Cut asm by simp | 
| 7167 |     also have "\<dots> = (Cut <a>.P (z).N){b:=(y).Q}" using Cut asm by (simp add: fresh_atm)
 | |
| 7168 |     also have "\<dots> = (Cut <a>.(P{b:=(y).Q}) (y).Q)" using Cut asm by (auto simp add: 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 | 7169 |     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 | 7170 |     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 | 7171 | using Cut asm by (simp add: fresh_atm) | 
| 7172 |     also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).(Q{x:=<a>.(P{b:=(y).Q})})" using Cut asm
 | |
| 36277 
9be4ab2acc13
split Class.thy into parts to 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 | by (auto simp add: fresh_prod fresh_atm subst_fresh) | 
| 41893 | 7174 |     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 | 7175 |     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 | 7176 | 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 | 7177 |     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 | 7178 | 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 | 7179 | } | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7181 |   { 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 | 7182 |     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 | 7183 | 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 | 7184 | 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 | 7185 | 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 | 7186 | 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 | 7187 | 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 | 7188 |     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 | 7189 | using Cut asm by simp | 
| 7190 |     also have "\<dots> = (Cut <c>.(M{x:=<a>.P}) (z).N){b:=(y).Q}" using Cut asm by (simp add: fresh_atm)
 | |
| 7191 |     also have "\<dots> = Cut <c>.(M{x:=<a>.P}{b:=(y).Q}) (y).Q" using Cut asm by (simp add: abs_fresh)
 | |
| 7192 |     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 | 7193 | 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 | 7194 |     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 | 7195 | 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 | 7196 |     have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
 | 
| 41893 | 7197 |                (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 | 7198 |     also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (y).(Q{x:=<a>.(P{b:=(y).Q})})"
 | 
| 41893 | 7199 | using Cut asm neq by (auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh) | 
| 7200 |     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 | 7201 |     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 | 7202 |                                        = 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 | 7203 |     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 | 7204 | 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 | 7205 | } | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7207 |   { 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 | 7208 |     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 | 7209 | 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 | 7210 | 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 | 7211 | 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 | 7212 | 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 | 7213 | 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 | 7214 |     have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (z).(N{x:=<a>.P})){b:=(y).Q}"
 | 
| 41893 | 7215 | using Cut asm by simp | 
| 7216 |     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 | 7217 | by (simp add: abs_fresh) | 
| 41893 | 7218 |     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 | 7219 |     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 | 7220 |                     = 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 | 7221 |     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 | 7222 |                     = (Cut <c>.(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}"
 | 
| 41893 | 7223 | 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 | 7224 |     also have "\<dots> = (Cut <c>.M (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}"
 | 
| 41893 | 7225 | using Cut asm by (auto simp add: 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 | 7226 |     also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" 
 | 
| 41893 | 7227 | 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 | 7228 | 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 | 7229 |     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 | 7230 |          = 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 | 7231 |     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 | 7232 | 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 | 7233 | } | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7235 |   { 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 | 7236 |     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 | 7237 | 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 | 7238 | 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 | 7239 | 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 | 7240 | 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 | 7241 | 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 | 7242 |     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 | 7243 | 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 | 7244 | 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 | 7245 | 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 | 7246 | 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 | 7247 | 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 | 7248 |     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 | 7249 | using Cut asm by simp | 
| 7250 |     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 | 7251 | 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 | 7252 |     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 | 7253 | 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 | 7254 |     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 | 7255 |              = 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 | 7256 |     have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
 | 
| 41893 | 7257 |                 (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 | 7258 |     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 | 7259 | 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 | 7260 |     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 | 7261 |            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 | 7262 |     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 | 7263 | 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 | 7264 | } | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7266 | 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 | 7267 | 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 | 7268 | 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 | 7269 | apply(auto 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 | 7270 |     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 | 7271 | 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 | 7272 | 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 | 7273 | 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 | 7274 | 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 | 7275 | 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 | 7276 | 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 | 7277 | 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 | 7278 | 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 | 7279 | 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 | 7280 | 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 | 7281 | 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 | 7282 | 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 | 7283 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7285 | 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 | 7286 | 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 | 7287 | apply(auto 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 | 7288 |     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 | 7289 | 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 | 7290 | 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 | 7291 | 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 | 7292 | 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 | 7293 | 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 | 7294 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7296 | 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 | 7297 | 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 | 7298 | apply(auto 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 | 7299 |     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 | 7300 |                                      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 | 7301 | 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 | 7302 | 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 | 7303 | 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 | 7304 | 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 | 7305 | 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 | 7306 | 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 | 7307 | 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 | 7308 | 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 | 7309 | 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 | 7310 | 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 | 7311 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7313 | 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 | 7314 | 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 | 7315 | apply(auto 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 | 7316 |     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 | 7317 | 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 | 7318 | 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 | 7319 | 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 | 7320 | 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 | 7321 | 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 | 7322 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7324 | 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 | 7325 | 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 | 7326 | apply(auto 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 | 7327 |     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 | 7328 | 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 | 7329 | 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 | 7330 | 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 | 7331 | 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 | 7332 | 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 | 7333 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7335 | 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 | 7336 | 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 | 7337 | apply(auto 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 | 7338 |     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 | 7339 |                                      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 | 7340 | 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 | 7341 | 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 | 7342 | 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 | 7343 | 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 | 7344 | 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 | 7345 | 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 | 7346 | 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 | 7347 | 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 | 7348 | 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 | 7349 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7350 | 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 | 7351 | 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 | 7352 | 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 | 7353 | apply(auto 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 | 7354 |     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 | 7355 |                                                      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 | 7356 | 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 | 7357 | 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 | 7358 | 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 | 7359 | 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 | 7360 | 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 | 7361 | 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 | 7362 | 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 | 7363 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7364 | 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 | 7365 | 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 | 7366 | 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 | 7367 | apply(auto 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 | 7368 |     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 | 7369 |                                                      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 | 7370 | 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 | 7371 | 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 | 7372 | 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 | 7373 | 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 | 7374 | 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 | 7375 | 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 | 7376 | 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 | 7377 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7378 | 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 | 7379 | 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 | 7380 | 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 | 7381 | apply(auto 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 | 7382 |     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 | 7383 |                                                      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 | 7384 | 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 | 7385 | 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 | 7386 | 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 | 7387 | 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 | 7388 | 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 | 7389 | 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 | 7390 | done | 
| 
9be4ab2acc13
split Class.thy into parts to 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 | 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 | 7392 | 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 | 7393 | 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 | 7394 | apply(auto 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 | 7395 |     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 | 7396 |                         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 | 7397 | 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 | 7398 | 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 | 7399 | 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 | 7400 | 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 | 7401 | 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 | 7402 | 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 | 7403 | 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 | 7404 | 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 | 7405 | 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 | 7406 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7407 | 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 | 7408 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7409 | 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 | 7410 | 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 | 7411 |   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 | 7412 | 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 | 7413 | 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 | 7414 | 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 | 7415 | 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 | 7416 | by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7417 | 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 | 7418 | 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 | 7419 | 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 | 7420 | 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 | 7421 | 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 | 7422 | 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 | 7423 | 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 | 7424 | 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 | 7425 | 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 | 7426 | 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 | 7427 | 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 | 7428 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7429 | 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 | 7430 | 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 | 7431 | 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 | 7432 | 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 | 7433 | 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 | 7434 | 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 | 7435 | 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 | 7436 | 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 | 7437 | 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 | 7438 | 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 | 7439 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7440 | 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 | 7441 | 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 | 7442 | 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 | 7443 | 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 | 7444 | 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 | 7445 | 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 | 7446 | 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 | 7447 | 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 | 7448 | 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 | 7449 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7450 | 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 | 7451 | 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 | 7452 | 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 | 7453 | 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 | 7454 | 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 | 7455 | 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 | 7456 | 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 | 7457 | 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 | 7458 | 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 | 7459 | 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 | 7460 | 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 | 7461 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7462 | 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 | 7463 | 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 | 7464 | 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 | 7465 | 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 | 7466 | 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 | 7467 | apply(auto)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7468 | 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 | 7469 | 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 | 7470 | 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 | 7471 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7472 | 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 | 7473 | 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 | 7474 | 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 | 7475 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7476 |     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 | 7477 | 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 | 7478 | 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 | 7479 | 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 | 7480 | 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 | 7481 | 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 | 7482 | 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 | 7483 | 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 | 7484 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7485 | 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 | 7486 | 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 | 7487 | 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 | 7488 | 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 | 7489 | 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 | 7490 | 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 | 7491 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7492 | 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 | 7493 | 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 | 7494 | 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 | 7495 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7496 |     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 | 7497 | 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 | 7498 | 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 | 7499 | 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 | 7500 | 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 | 7501 | 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 | 7502 | 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 | 7503 | 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 | 7504 | 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 | 7505 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7506 | 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 | 7507 | 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 | 7508 | 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 | 7509 | 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 | 7510 | 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 | 7511 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7512 | 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 | 7513 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7514 | 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 | 7515 | 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 | 7516 | 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 | 7517 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7518 |     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 | 7519 |                   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 | 7520 | 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 | 7521 | 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 | 7522 | 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 | 7523 | 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 | 7524 | 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 | 7525 | 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 | 7526 | 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 | 7527 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7528 | 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 | 7529 | 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 | 7530 | 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 | 7531 | 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 | 7532 | 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 | 7533 | 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 | 7534 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7535 | 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 | 7536 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7537 | 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 | 7538 | 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 | 7539 | 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 | 7540 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7541 |     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 | 7542 | 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 | 7543 | 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 | 7544 | 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 | 7545 | 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 | 7546 | 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 | 7547 | 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 | 7548 | 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 | 7549 | 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 | 7550 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7551 | 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 | 7552 | 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 | 7553 | 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 | 7554 | 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 | 7555 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7556 | 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 | 7557 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7558 | 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 | 7559 | 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 | 7560 | 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 | 7561 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7562 |     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 | 7563 | 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 | 7564 | 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 | 7565 | 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 | 7566 | 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 | 7567 | 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 | 7568 | 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 | 7569 | 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 | 7570 | 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 | 7571 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7572 | 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 | 7573 | 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 | 7574 | 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 | 7575 | 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 | 7576 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7577 | 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 | 7578 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7579 | 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 | 7580 | 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 | 7581 | 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 | 7582 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7583 |     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 | 7584 |                   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 | 7585 | 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 | 7586 | 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 | 7587 | 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 | 7588 | 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 | 7589 | 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 | 7590 | 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 | 7591 | 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 | 7592 | 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 | 7593 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7594 | 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 | 7595 | 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 | 7596 | 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 | 7597 | 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 | 7598 | 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 | 7599 | 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 | 7600 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7601 | 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 | 7602 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7603 | 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 | 7604 | 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 | 7605 | 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 | 7606 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7607 |     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 | 7608 |                                         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 | 7609 | 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 | 7610 | 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 | 7611 | 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 | 7612 | 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 | 7613 | 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 | 7614 | 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 | 7615 | 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 | 7616 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7617 | 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 | 7618 | 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 | 7619 | 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 | 7620 | 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 | 7621 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7622 | 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 | 7623 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7624 | 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 | 7625 | 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 | 7626 | 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 | 7627 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7628 |     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 | 7629 |                                         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 | 7630 | 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 | 7631 | 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 | 7632 | 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 | 7633 | 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 | 7634 | 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 | 7635 | 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 | 7636 | 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 | 7637 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7638 | 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 | 7639 | 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 | 7640 | 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 | 7641 | 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 | 7642 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7643 | 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 | 7644 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7645 | 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 | 7646 | 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 | 7647 | 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 | 7648 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7649 |     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 | 7650 |                                         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 | 7651 | 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 | 7652 | 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 | 7653 | 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 | 7654 | 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 | 7655 | 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 | 7656 | 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 | 7657 | 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 | 7658 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7659 | 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 | 7660 | 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 | 7661 | 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 | 7662 | 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 | 7663 | 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 | 7664 | 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 | 7665 | 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 | 7666 | 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 | 7667 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7668 | 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 | 7669 | 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 | 7670 | 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 | 7671 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget 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 | 7672 |     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 | 7673 |                   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 | 7674 | 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 | 7675 | 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 | 7676 | 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 | 7677 | 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 | 7678 | 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 | 7679 | 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 | 7680 | 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 | 7681 | 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 | 7682 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7683 | 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 | 7684 | 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 | 7685 | 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 | 7686 | 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 | 7687 | 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 | 7688 | 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 | 7689 | 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 | 7690 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7691 | 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 | 7692 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7693 | 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 | 7694 | 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 | 7695 |   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 | 7696 | 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 | 7697 | 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 | 7698 | 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 | 7699 | 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 | 7700 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7701 | 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 | 7702 | 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 | 7703 | 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 | 7704 | 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 | 7705 | 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 | 7706 | 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 | 7707 | 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 | 7708 | 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 | 7709 | 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 | 7710 | 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 | 7711 | 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 | 7712 | 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 | 7713 | 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 | 7714 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7715 | 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 | 7716 | 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 | 7717 | 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 | 7718 | 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 | 7719 | 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 | 7720 | 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 | 7721 | 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 | 7722 | 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 | 7723 | 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 | 7724 | 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 | 7725 | 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 | 7726 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7727 | 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 | 7728 | 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 | 7729 | 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 | 7730 | 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 | 7731 | 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 | 7732 | 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 | 7733 | 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 | 7734 | 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 | 7735 | 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 | 7736 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7737 | 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 | 7738 | 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 | 7739 | 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 | 7740 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7741 | 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 | 7742 | 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 | 7743 | 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 | 7744 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7745 |     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 | 7746 | 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 | 7747 | 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 | 7748 | 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 | 7749 | 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 | 7750 | 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 | 7751 | 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 | 7752 | 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 | 7753 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7754 | 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 | 7755 | 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 | 7756 | 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 | 7757 | 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 | 7758 | 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 | 7759 | 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 | 7760 |     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 | 7761 | 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 | 7762 | 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 | 7763 | 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 | 7764 | 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 | 7765 | 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 | 7766 | 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 | 7767 | 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 | 7768 | 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 | 7769 | 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 | 7770 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7771 | 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 | 7772 | 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 | 7773 | 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 | 7774 | 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 | 7775 | 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 | 7776 | 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 | 7777 | 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 | 7778 | 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 | 7779 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7780 | 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 | 7781 | 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 | 7782 | 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 | 7783 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7784 | 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 | 7785 | 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 | 7786 | 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 | 7787 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7788 |     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 | 7789 | 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 | 7790 | 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 | 7791 | 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 | 7792 | 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 | 7793 | 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 | 7794 | 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 | 7795 | 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 | 7796 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7797 | 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 | 7798 | 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 | 7799 | 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 | 7800 | 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 | 7801 | 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 | 7802 | 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 | 7803 |     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 | 7804 | 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 | 7805 | 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 | 7806 | 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 | 7807 | 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 | 7808 | 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 | 7809 | 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 | 7810 | 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 | 7811 | 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 | 7812 | 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 | 7813 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7814 | 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 | 7815 | 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 | 7816 | 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 | 7817 | 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 | 7818 | 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 | 7819 | 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 | 7820 | 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 | 7821 | 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 | 7822 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7823 | 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 | 7824 | 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 | 7825 | 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 | 7826 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7827 |     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 | 7828 | 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 | 7829 | 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 | 7830 | 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 | 7831 | 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 | 7832 | 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 | 7833 | 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 | 7834 | 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 | 7835 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7836 | 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 | 7837 | 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 | 7838 | 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 | 7839 | 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 | 7840 | 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 | 7841 | 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 | 7842 |     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 | 7843 | 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 | 7844 | 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 | 7845 | 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 | 7846 | 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 | 7847 | 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 | 7848 | 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 | 7849 | 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 | 7850 | 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 | 7851 | 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 | 7852 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7853 | 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 | 7854 | 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 | 7855 | 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 | 7856 | 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 | 7857 | 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 | 7858 | 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 | 7859 | 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 | 7860 | 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 | 7861 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7862 | 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 | 7863 | 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 | 7864 | 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 | 7865 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7866 | 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 | 7867 | 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 | 7868 | 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 | 7869 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7870 | 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 | 7871 | 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 | 7872 | 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 | 7873 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7874 |     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 | 7875 |                                       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 | 7876 | 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 | 7877 | 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 | 7878 | 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 | 7879 | 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 | 7880 | 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 | 7881 | 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 | 7882 | 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 | 7883 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7884 | 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 | 7885 | 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 | 7886 | 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 | 7887 | 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 | 7888 | 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 | 7889 | 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 | 7890 | 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 | 7891 | 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 | 7892 |     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 | 7893 |                                       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 | 7894 | 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 | 7895 | 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 | 7896 | 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 | 7897 | 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 | 7898 | 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 | 7899 | 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 | 7900 | 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 | 7901 | 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 | 7902 | 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 | 7903 | 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 | 7904 | 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 | 7905 | 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 | 7906 | 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 | 7907 | 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 | 7908 | 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 | 7909 | 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 | 7910 | 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 | 7911 | 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 | 7912 | 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 | 7913 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7914 | 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 | 7915 | 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 | 7916 | 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 | 7917 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7918 | 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 | 7919 | 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 | 7920 | 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 | 7921 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7922 |     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 | 7923 |                                       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 | 7924 | 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 | 7925 | 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 | 7926 | 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 | 7927 | 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 | 7928 | 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 | 7929 | 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 | 7930 | 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 | 7931 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7932 | 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 | 7933 | 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 | 7934 | 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 | 7935 | 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 | 7936 | 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 | 7937 | 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 | 7938 | 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 | 7939 |     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 | 7940 |                                       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 | 7941 | 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 | 7942 | 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 | 7943 | 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 | 7944 | 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 | 7945 | 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 | 7946 | 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 | 7947 | 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 | 7948 | 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 | 7949 | 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 | 7950 | 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 | 7951 | 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 | 7952 | 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 | 7953 | 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 | 7954 | 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 | 7955 | 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 | 7956 | 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 | 7957 | 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 | 7958 | 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 | 7959 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7960 | 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 | 7961 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7962 | 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 | 7963 | 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 | 7964 |   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 | 7965 | 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 | 7966 | 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 | 7967 | 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 | 7968 | 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 | 7969 | by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 7970 | 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 | 7971 | 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 | 7972 | 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 | 7973 | 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 | 7974 | 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 | 7975 | 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 | 7976 | 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 | 7977 | 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 | 7978 | 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 | 7979 | 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 | 7980 | 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 | 7981 | 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 | 7982 | 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 | 7983 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7984 | 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 | 7985 | 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 | 7986 | 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 | 7987 | 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 | 7988 | 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 | 7989 | 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 | 7990 | 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 | 7991 | 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 | 7992 | 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 | 7993 | 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 | 7994 | 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 | 7995 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 7996 | 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 | 7997 | 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 | 7998 | 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 | 7999 | 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 | 8000 | 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 | 8001 | 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 | 8002 | 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 | 8003 | 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 | 8004 | 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 | 8005 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8006 | 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 | 8007 | 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 | 8008 | 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 | 8009 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8010 | 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 | 8011 | 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 | 8012 | 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 | 8013 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8014 | 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 | 8015 | 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 | 8016 | 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 | 8017 | 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 | 8018 | 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 | 8019 | 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 | 8020 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8021 | 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 | 8022 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8023 | 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 | 8024 | 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 | 8025 | 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 | 8026 | 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 | 8027 | apply(auto simp add: fresh_atm fresh_prod)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8028 | 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 | 8029 | 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 | 8030 | 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 | 8031 | 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 | 8032 | 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 | 8033 | 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 | 8034 | 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 | 8035 | 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 | 8036 | apply(auto 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 | 8037 | 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 | 8038 | 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 | 8039 | 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 | 8040 | 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 | 8041 | apply(auto 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 | 8042 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8043 | 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 | 8044 | 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 | 8045 | 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 | 8046 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8047 | 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 | 8048 | 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 | 8049 | 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 | 8050 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8051 | 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 | 8052 | 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 | 8053 | 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 | 8054 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8055 | 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 | 8056 | 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 | 8057 | 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 | 8058 | 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 | 8059 | 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 | 8060 | 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 | 8061 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8062 | 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 | 8063 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8064 | 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 | 8065 | 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 | 8066 | 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 | 8067 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8068 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8069 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8070 | apply(auto simp add: fresh_atm fresh_prod)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8071 | 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 | 8072 | 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 | 8073 | 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 | 8074 | 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 | 8075 | 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 | 8076 | 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 | 8077 | 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 | 8078 | 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 | 8079 | apply(auto simp add: fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8080 | 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 | 8081 | 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 | 8082 | 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 | 8083 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8084 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8085 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8086 | apply(auto simp add: fresh_atm fresh_prod)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8087 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8088 | 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 | 8089 | 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 | 8090 | 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 | 8091 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8092 | 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 | 8093 | 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 | 8094 | 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 | 8095 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8096 | 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 | 8097 | 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 | 8098 | 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 | 8099 | 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 | 8100 | 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 | 8101 | 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 | 8102 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8103 | 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 | 8104 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8105 | 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 | 8106 | 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 | 8107 | 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 | 8108 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8109 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8110 | 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 | 8111 | 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 | 8112 | 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 | 8113 | 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 | 8114 | 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 | 8115 | 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 | 8116 | 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 | 8117 | 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 | 8118 | apply(auto simp add: fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8119 | 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 | 8120 | 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 | 8121 | 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 | 8122 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8123 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8124 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8125 | 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 | 8126 | 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 | 8127 | 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 | 8128 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8129 | 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 | 8130 | 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 | 8131 | 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 | 8132 | 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 | 8133 | 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 | 8134 | 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 | 8135 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8136 | 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 | 8137 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8138 | 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 | 8139 | 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 | 8140 | 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 | 8141 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8142 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8143 | 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 | 8144 | 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 | 8145 | 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 | 8146 | 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 | 8147 | 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 | 8148 | 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 | 8149 | 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 | 8150 | 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 | 8151 | apply(auto simp add: fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8152 | 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 | 8153 | 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 | 8154 | 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 | 8155 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8156 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8157 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8158 | 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 | 8159 | 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 | 8160 | 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 | 8161 | by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8162 | 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 | 8163 | 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 | 8164 | 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 | 8165 | apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod 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 | 8166 | 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 | 8167 | 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 | 8168 | 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 | 8169 | 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 | 8170 | 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 | 8171 | 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 | 8172 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8173 | 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 | 8174 | apply(simp) | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8175 | 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 | 8176 | 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 | 8177 | 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 | 8178 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8179 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8180 | apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8181 | 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 | 8182 | 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 | 8183 | 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 | 8184 | 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 | 8185 | 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 | 8186 | 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 | 8187 | 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 | 8188 | 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 | 8189 | apply(auto simp add: fresh_atm)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8190 | 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 | 8191 | 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 | 8192 | 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 | 8193 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8194 | apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8195 | apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8196 | apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8197 | done | 
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8198 | 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 | 8199 | |
| 
9be4ab2acc13
split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
 wenzelm parents: diff
changeset | 8200 | end |