src/HOL/Nominal/Examples/Class1.thy
author wenzelm
Tue, 08 Oct 2024 16:15:31 +0200
changeset 81132 dff7dfd8dce3
parent 80914 d97fdabd9e2b
permissions -rw-r--r--
more robust declarations via "no syntax" bundles;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     1
theory Class1
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
     2
imports "HOL-Nominal.Nominal"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     3
begin
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     4
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
     5
section \<open>Term-Calculus from Urban's PhD\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     6
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     7
atom_decl name coname
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
     8
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
     9
text \<open>types\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    10
81132
dff7dfd8dce3 more robust declarations via "no syntax" bundles;
wenzelm
parents: 80914
diff changeset
    11
unbundle no bit_operations_syntax
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73932
diff changeset
    12
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    13
nominal_datatype ty =
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    14
    PR "string"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    15
  | NOT  "ty"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    16
  | AND  "ty" "ty"   (\<open>_ AND _\<close> [100,100] 100)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    17
  | OR   "ty" "ty"   (\<open>_ OR _\<close>  [100,100] 100)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    18
  | IMP  "ty" "ty"   (\<open>_ IMP _\<close> [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    19
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    20
instantiation ty :: size
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    21
begin
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    22
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    23
nominal_primrec size_ty
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    24
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    25
  "size (PR s)     = (1::nat)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    26
| "size (NOT T)     = 1 + size T"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    27
| "size (T1 AND T2) = 1 + size T1 + size T2"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    28
| "size (T1 OR T2)  = 1 + size T1 + size T2"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    29
| "size (T1 IMP T2) = 1 + size T1 + size T2"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    30
by (rule TrueI)+
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    31
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    32
instance ..
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    33
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    34
end
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    35
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    36
lemma ty_cases:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    37
  fixes T::ty
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    38
  shows "(\<exists>s. T=PR s) \<or> (\<exists>T'. T=NOT T') \<or> (\<exists>S U. T=S OR U) \<or> (\<exists>S U. T=S AND U) \<or> (\<exists>S U. T=S IMP U)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    39
by (induct T rule:ty.induct) (auto)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    40
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    41
lemma fresh_ty:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    42
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    43
  and   x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    44
  and   T::"ty"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    45
  shows "a\<sharp>T" and "x\<sharp>T"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    46
by (nominal_induct T rule: ty.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    47
   (auto simp: fresh_string)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    48
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    49
text \<open>terms\<close>                               
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    50
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    51
nominal_datatype trm = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    52
    Ax   "name" "coname"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    53
  | Cut  "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"            (\<open>Cut <_>._ ('(_'))._\<close> [0,0,0,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    54
  | NotR "\<guillemotleft>name\<guillemotright>trm" "coname"                 (\<open>NotR ('(_'))._ _\<close> [0,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    55
  | NotL "\<guillemotleft>coname\<guillemotright>trm" "name"                 (\<open>NotL <_>._ _\<close> [0,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    56
  | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" (\<open>AndR <_>._ <_>._ _\<close> [0,100,0,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    57
  | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name"                  (\<open>AndL1 (_)._ _\<close> [100,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    58
  | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name"                  (\<open>AndL2 (_)._ _\<close> [100,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    59
  | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname"               (\<open>OrR1 <_>._ _\<close> [100,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    60
  | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname"               (\<open>OrR2 <_>._ _\<close> [100,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    61
  | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"        (\<open>OrL (_)._ (_)._ _\<close> [100,100,100,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    62
  | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"       (\<open>ImpR (_).<_>._ _\<close> [100,100,100,100] 101)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    63
  | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"     (\<open>ImpL <_>._ (_)._ _\<close> [100,100,100,100,100] 101)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    64
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
    65
text \<open>named terms\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    66
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    67
nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" (\<open>((_):_)\<close> [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    68
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
    69
text \<open>conamed terms\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    70
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    71
nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" (\<open>(<_>:_)\<close> [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    72
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
    73
text \<open>renaming functions\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    74
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    75
nominal_primrec (freshness_context: "(d::coname,e::coname)") 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    76
  crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm"  (\<open>_[_\<turnstile>c>_]\<close> [100,0,0] 100) 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    77
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    78
  "(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    79
| "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>M\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>e])" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    80
| "(NotR (x).M a)[d\<turnstile>c>e] = (if a=d then NotR (x).(M[d\<turnstile>c>e]) e else NotR (x).(M[d\<turnstile>c>e]) a)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    81
| "a\<sharp>(d,e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    82
| "\<lbrakk>a\<sharp>(d,e,N,c);b\<sharp>(d,e,M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>e] = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    83
          (if c=d then AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d \<turnstile>c>e]) e else AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d\<turnstile>c>e]) c)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    84
| "x\<sharp>y \<Longrightarrow> (AndL1 (x).M y)[d\<turnstile>c>e] = AndL1 (x).(M[d\<turnstile>c>e]) y"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    85
| "x\<sharp>y \<Longrightarrow> (AndL2 (x).M y)[d\<turnstile>c>e] = AndL2 (x).(M[d\<turnstile>c>e]) y"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    86
| "a\<sharp>(d,e,b) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    87
| "a\<sharp>(d,e,b) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    88
| "\<lbrakk>x\<sharp>(N,z);y\<sharp>(M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N z)[d\<turnstile>c>e] = OrL (x).(M[d\<turnstile>c>e]) (y).(N[d\<turnstile>c>e]) z"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    89
| "a\<sharp>(d,e,b) \<Longrightarrow> (ImpR (x).<a>.M b)[d\<turnstile>c>e] = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    90
       (if b=d then ImpR (x).<a>.(M[d\<turnstile>c>e]) e else ImpR (x).<a>.(M[d\<turnstile>c>e]) b)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    91
| "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>(M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y)[d\<turnstile>c>e] = ImpL <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>e]) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
    92
  by(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    93
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    94
nominal_primrec (freshness_context: "(u::name,v::name)") 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
    95
  nrename :: "trm \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm"      (\<open>_[_\<turnstile>n>_]\<close> [100,0,0] 100) 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    96
where
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    97
  "(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else Ax x a)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    98
| "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
    99
| "x\<sharp>(u,v) \<Longrightarrow> (NotR (x).M a)[u\<turnstile>n>v] = NotR (x).(M[u\<turnstile>n>v]) a" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   100
| "(NotL <a>.M x)[u\<turnstile>n>v] = (if x=u then NotL <a>.(M[u\<turnstile>n>v]) v else NotL <a>.(M[u\<turnstile>n>v]) x)" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   101
| "\<lbrakk>a\<sharp>(N,c);b\<sharp>(M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[u\<turnstile>n>v] = AndR <a>.(M[u\<turnstile>n>v]) <b>.(N[u\<turnstile>n>v]) c" 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   102
| "x\<sharp>(u,v,y) \<Longrightarrow> (AndL1 (x).M y)[u\<turnstile>n>v] = (if y=u then AndL1 (x).(M[u\<turnstile>n>v]) v else AndL1 (x).(M[u\<turnstile>n>v]) y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   103
| "x\<sharp>(u,v,y) \<Longrightarrow> (AndL2 (x).M y)[u\<turnstile>n>v] = (if y=u then AndL2 (x).(M[u\<turnstile>n>v]) v else AndL2 (x).(M[u\<turnstile>n>v]) y)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   104
| "a\<sharp>b \<Longrightarrow> (OrR1 <a>.M b)[u\<turnstile>n>v] = OrR1 <a>.(M[u\<turnstile>n>v]) b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   105
| "a\<sharp>b \<Longrightarrow> (OrR2 <a>.M b)[u\<turnstile>n>v] = OrR2 <a>.(M[u\<turnstile>n>v]) b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   106
| "\<lbrakk>x\<sharp>(u,v,N,z);y\<sharp>(u,v,M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N z)[u\<turnstile>n>v] = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   107
        (if z=u then OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) v else OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) z)"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   108
| "\<lbrakk>a\<sharp>b; x\<sharp>(u,v)\<rbrakk> \<Longrightarrow> (ImpR (x).<a>.M b)[u\<turnstile>n>v] = ImpR (x).<a>.(M[u\<turnstile>n>v]) b"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   109
| "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y)[u\<turnstile>n>v] = 
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   110
        (if y=u then ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) v else ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) y)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   111
  by(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   112
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   113
lemmas eq_bij = pt_bij[OF pt_name_inst, OF at_name_inst] pt_bij[OF pt_coname_inst, OF at_coname_inst]
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   114
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   115
lemma crename_name_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   116
  fixes pi::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   117
  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   118
  by (nominal_induct M avoiding: d e rule: trm.strong_induct) (auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   119
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   120
lemma crename_coname_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   121
  fixes pi::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   122
  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   123
  by (nominal_induct M avoiding: d e rule: trm.strong_induct)(auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   124
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   125
lemma nrename_name_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   126
  fixes pi::"name prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   127
  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   128
  by(nominal_induct M avoiding: x y rule: trm.strong_induct) (auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   129
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   130
lemma nrename_coname_eqvt[eqvt]:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   131
  fixes pi::"coname prm"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   132
  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   133
  by(nominal_induct M avoiding: x y rule: trm.strong_induct)(auto simp: fresh_bij eq_bij)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   134
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   135
lemmas rename_eqvts = crename_name_eqvt crename_coname_eqvt
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   136
                      nrename_name_eqvt nrename_coname_eqvt
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   137
lemma nrename_fresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   138
  assumes a: "x\<sharp>M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   139
  shows "M[x\<turnstile>n>y] = M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   140
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   141
by (nominal_induct M avoiding: x y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   142
   (auto simp: trm.inject fresh_atm abs_fresh fin_supp abs_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   143
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   144
lemma crename_fresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   145
  assumes a: "a\<sharp>M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   146
  shows "M[a\<turnstile>c>b] = M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   147
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   148
by (nominal_induct M avoiding: a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   149
   (auto simp: trm.inject fresh_atm abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   150
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   151
lemma nrename_nfresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   152
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   153
  shows "x\<sharp>y\<Longrightarrow>x\<sharp>M[x\<turnstile>n>y]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   154
by (nominal_induct M avoiding: x y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   155
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   156
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   157
 lemma crename_nfresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   158
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   159
  shows "x\<sharp>M\<Longrightarrow>x\<sharp>M[a\<turnstile>c>b]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   160
by (nominal_induct M avoiding: a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   161
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   162
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   163
lemma crename_cfresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   164
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   165
  shows "a\<sharp>b\<Longrightarrow>a\<sharp>M[a\<turnstile>c>b]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   166
by (nominal_induct M avoiding: a b rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   167
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   168
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   169
lemma nrename_cfresh:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   170
  fixes c::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   171
  shows "c\<sharp>M\<Longrightarrow>c\<sharp>M[x\<turnstile>n>y]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   172
by (nominal_induct M avoiding: x y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   173
   (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   174
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   175
lemma nrename_nfresh':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   176
  fixes x::"name"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   177
  shows "x\<sharp>(M,z,y)\<Longrightarrow>x\<sharp>M[z\<turnstile>n>y]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   178
by (nominal_induct M avoiding: x z y rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   179
   (auto simp: fresh_prod fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   180
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   181
lemma crename_cfresh':
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   182
  fixes a::"coname"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   183
  shows "a\<sharp>(M,b,c)\<Longrightarrow>a\<sharp>M[b\<turnstile>c>c]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   184
by (nominal_induct M avoiding: a b c rule: trm.strong_induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   185
   (auto simp: fresh_prod fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   186
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   187
lemma nrename_rename:
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   188
  assumes a: "x'\<sharp>M"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   189
  shows "([(x',x)]\<bullet>M)[x'\<turnstile>n>y]= M[x\<turnstile>n>y]"
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   190
using a
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   191
proof (nominal_induct M avoiding: x x' y rule: trm.strong_induct)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   192
qed (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   193
9be4ab2acc13 split Class.thy into parts to 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
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
   195
  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
   196
  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
   197
using a
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   198
proof (nominal_induct M avoiding: a a' b rule: trm.strong_induct)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   199
qed (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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
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
   202
                      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
   203
                      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
   204
                      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
   205
9be4ab2acc13 split Class.thy into parts to 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
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
   207
  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
   208
  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
   209
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
   210
  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
   211
  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
   212
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   213
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   214
  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
   215
                        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
   216
    using fs1 fs2
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   217
    by (intro nrename.simps; simp add: fresh_left calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   218
  also have "\<dots> = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])" using fs1 fs2 a
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   219
    by(simp add: trm.inject alpha fresh_prod rename_eqvts calc_atm rename_fresh fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   220
  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
   221
    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
   222
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
   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_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
   225
  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
   226
  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
   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,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
   229
  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
   230
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   231
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   232
  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
   233
                        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
   234
    using fs1 fs2
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   235
    by (intro crename.simps; simp add: fresh_left calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   236
  also have "\<dots> = Cut <a>.(M[b\<turnstile>c>c]) (x).(N[b\<turnstile>c>c])" using fs1 fs2 a
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   237
    by(simp add: trm.inject alpha calc_atm rename_fresh fresh_atm fresh_prod rename_eqvts)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   238
  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
   239
    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
   240
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
   241
9be4ab2acc13 split Class.thy into parts to 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
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
   243
  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
   244
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
   245
9be4ab2acc13 split Class.thy into parts to 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
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
   247
  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
   248
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
   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 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
   251
  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
   252
by (nominal_induct M avoiding: x y rule: trm.strong_induct) 
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   253
   (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   254
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
   257
  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
   258
by (nominal_induct M avoiding: a b rule: trm.strong_induct) 
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   259
   (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   260
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
   263
  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
   264
  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
   265
using a
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   266
proof (nominal_induct M avoiding: a b x c rule: trm.strong_induct)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   267
qed (simp_all add: trm.inject split: if_splits)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
   270
  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
   271
  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
   272
using a
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   273
proof (nominal_induct M avoiding: x y z a rule: trm.strong_induct)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   274
qed (simp_all add: trm.inject split: if_splits)
36277
9be4ab2acc13 split Class.thy into parts to 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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
   276
text \<open>substitution functions\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   277
9be4ab2acc13 split Class.thy into parts to 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
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
   279
  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
   280
  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
   281
  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
   282
  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
   283
  shows "c\<sharp>(pi\<bullet>M)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   284
  by (simp add: assms fresh_perm_app)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   285
9be4ab2acc13 split Class.thy into parts to 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
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
   287
  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
   288
  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
   289
  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
   290
  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
   291
  shows "x\<sharp>(pi\<bullet>M)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   292
  by (simp add: assms fresh_perm_app)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   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 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
   295
  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
   296
  shows "fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x') = Cut <c>.P (x').NotL <a>.M x'"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   297
proof (rule fresh_fun_app)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   298
  show "pt (TYPE(trm)::trm itself) (TYPE(name)::name itself)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   299
    by(rule pt_name_inst)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   300
  show "at (TYPE(name)::name itself)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   301
    by(rule at_name_inst)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   302
  show "finite (supp (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')::name set)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   303
    using a by(finite_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   304
  obtain n::name where n: "n\<sharp>(c,P,a,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   305
    by (metis assms fresh_atm(3) fresh_prod)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   306
  with assms have "n \<sharp> (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   307
    by (fresh_guess)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   308
  then show "\<exists>b. b \<sharp> (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x', Cut <c>.P (b).NotL <a>.M b)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   309
    by (metis abs_fresh(1) abs_fresh(5) fresh_prod n trm.fresh(3))
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   310
  show "x' \<sharp> (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   311
    using assms by(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   312
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   313
9be4ab2acc13 split Class.thy into parts to 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
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
   315
  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
   316
  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
   317
  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
   318
             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
   319
  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
   320
             fresh_fun (pi2\<bullet>(\<lambda>x'. Cut <c>.P (x').NotL <a>.M x'))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   321
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   322
   apply(generate_fresh "name")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   323
   apply(auto simp: fresh_prod fresh_fun_simp_NotL)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
   324
  apply (metis fresh_bij(1) fresh_fun_simp_NotL name_prm_coname_def)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   325
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   326
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   327
   apply (metis fresh_fun_simp_NotL fresh_prodD swap_simps(8) trm.perm(14) trm.perm(16))
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   328
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   329
9be4ab2acc13 split Class.thy into parts to 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
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
   331
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>x"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   332
  shows "fresh_fun (\<lambda>z'. Cut <c>.P(z').AndL1 (x).M z') = Cut <c>.P (z').AndL1 (x).M z'"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   333
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   334
  obtain n::name where "n\<sharp>(c,P,x,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   335
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   336
  then show "\<exists>a. a \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL1 x. M z', Cut <c>.P(a).AndL1 x. M a)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   337
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   338
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   339
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   340
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   341
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   342
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL1 x. M z')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   343
    using a by(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   344
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   345
9be4ab2acc13 split Class.thy into parts to 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
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
   347
  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
   348
  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
   349
  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
   350
             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
   351
  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
   352
             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
   353
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
   354
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   355
apply(force simp add: fresh_prod fresh_fun_simp_AndL1 at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   356
  apply (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   357
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
   358
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)")
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   359
apply(force simp add: fresh_prod fresh_fun_simp_AndL1 calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   360
  by (meson exists_fresh'(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   361
9be4ab2acc13 split Class.thy into parts to 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
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
   363
  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
   364
  shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z') = Cut <c>.P (z').AndL2 (x).M z'"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   365
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   366
  obtain n::name where "n\<sharp>(c,P,x,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   367
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   368
  then show "\<exists>a. a \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL2 x. M z', Cut <c>.P(a).AndL2 x. M a)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   369
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   370
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   371
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   372
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   373
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   374
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').AndL2 x. M z')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   375
    using a by(fresh_guess)           
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   376
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   377
9be4ab2acc13 split Class.thy into parts to 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
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
   379
  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
   380
  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
   381
  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
   382
             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
   383
  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
   384
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z'))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   385
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   386
   apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   387
    apply(force simp add: fresh_prod fresh_fun_simp_AndL2 at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   388
   apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   389
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   390
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   391
   apply(force simp add: fresh_prod fresh_fun_simp_AndL2 calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   392
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   393
9be4ab2acc13 split Class.thy into parts to 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
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
   395
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>N" "z'\<sharp>u" "z'\<sharp>x"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   396
  shows "fresh_fun (\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N z') = Cut <c>.P (z').OrL (x).M (u).N z'"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   397
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   398
  obtain n::name where "n\<sharp>(c,P,x,M,u,N)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   399
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   400
  then show "\<exists>a. a \<sharp> (\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N(z'), Cut <c>.P(a).OrL(x).M(u).N(a))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   401
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   402
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   403
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   404
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   405
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   406
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N(z'))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   407
    using a by(fresh_guess) 
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   408
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   409
9be4ab2acc13 split Class.thy into parts to 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
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
   411
  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
   412
  and   pi2::"coname prm"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   413
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P(z').OrL (x).M(u).N z')=
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   414
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').OrL(x).M (u).N z'))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   415
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P(z').OrL (x).M(u).N z')=
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   416
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P(z').OrL(x).M(u).N z'))"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   417
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   418
   apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,u,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>N,pi1\<bullet>x,pi1\<bullet>u,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   419
    apply(force simp add: fresh_prod fresh_fun_simp_OrL at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   420
   apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   421
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   422
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,u,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>N,pi2\<bullet>x,pi2\<bullet>u,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   423
   apply(force simp add: fresh_prod fresh_fun_simp_OrL calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   424
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   425
9be4ab2acc13 split Class.thy into parts to 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
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
   427
  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
   428
  shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z') = Cut <c>.P (z').ImpL <a>.M (x).N z'"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   429
proof (rule fresh_fun_app [OF pt_name_inst at_name_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   430
  obtain n::name where "n\<sharp>(c,P,x,M,N)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   431
    by (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   432
  then show "\<exists>aa. aa \<sharp> (\<lambda>z'. Cut <c>.P(z').ImpL <a>.M(x).N z', Cut <c>.P(aa).ImpL <a>.M(x).N aa)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   433
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   434
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   435
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   436
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   437
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   438
  show "z' \<sharp> (\<lambda>z'. Cut <c>.P(z').ImpL <a>.M(x).N z')"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   439
    using a by(fresh_guess) 
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   440
qed finite_guess
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   441
9be4ab2acc13 split Class.thy into parts to 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
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
   443
  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
   444
  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
   445
  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
   446
             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
   447
  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
   448
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z'))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   449
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   450
   apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>N,pi1\<bullet>x,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   451
    apply(force simp add: fresh_prod fresh_fun_simp_ImpL at_prm_fresh[OF at_name_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   452
   apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   453
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   454
  apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,N,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>N,pi2\<bullet>x,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   455
   apply(force simp add: fresh_prod fresh_fun_simp_ImpL calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   456
  by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   457
9be4ab2acc13 split Class.thy into parts to 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
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
   459
  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
   460
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P) = Cut <a'>.(NotR (y).M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   461
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   462
  obtain n::coname where "n\<sharp>(x,P,y,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   463
    by (metis assms(1) assms(2) fresh_atm(4) fresh_prod)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   464
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P, Cut <a>.NotR(y).M(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   465
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   466
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   467
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   468
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   469
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   470
9be4ab2acc13 split Class.thy into parts to 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
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
   472
  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
   473
  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
   474
  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
   475
             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
   476
  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
   477
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   478
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   479
   apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi1\<bullet>P,pi1\<bullet>M,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   480
    apply(force simp add: fresh_prod fresh_fun_simp_NotR calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   481
   apply (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   482
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   483
  apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   484
   apply(force simp: fresh_prod fresh_fun_simp_NotR at_prm_fresh[OF at_coname_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   485
  by (meson exists_fresh(2) fs_coname1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   486
9be4ab2acc13 split Class.thy into parts to 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
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
   488
  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
   489
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P) = Cut <a'>.(AndR <b>.M <c>.N a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   490
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   491
  obtain n::coname where "n\<sharp>(x,P,b,M,c,N)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   492
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   493
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.AndR <b>.M <c>.N(a') (x).P, Cut <a>.AndR <b>.M <c>.N(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   494
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   495
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   496
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   497
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   498
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   499
9be4ab2acc13 split Class.thy into parts to 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
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
   501
  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
   502
  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
   503
  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
   504
             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
   505
  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
   506
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   507
   apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   508
   apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,N,b,c,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>N,pi1\<bullet>b,pi1\<bullet>c,pi1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   509
    apply(force simp add: fresh_prod fresh_fun_simp_AndR calc_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   510
   apply (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   511
  apply(perm_simp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   512
  apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,N,b,c,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>N,pi2\<bullet>b,pi2\<bullet>c,pi2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   513
   apply(force simp add: fresh_prod fresh_fun_simp_AndR at_prm_fresh[OF at_coname_inst] swap_simps)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   514
  by (meson exists_fresh(2) fs_coname1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   515
9be4ab2acc13 split Class.thy into parts to 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
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
   517
  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
   518
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P) = Cut <a'>.(OrR1 <b>.M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   519
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   520
  obtain n::coname where "n\<sharp>(x,P,b,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   521
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   522
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.OrR1 <b>.M(a') (x).P, Cut <a>.OrR1 <b>.M(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   523
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   524
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   525
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   526
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   527
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   528
9be4ab2acc13 split Class.thy into parts to 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
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
   530
  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
   531
  and   pi2::"coname prm"
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   532
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P) =
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   533
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(OrR1 <b>.M  a') (x).P))" (is "?t1")
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   534
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P) =
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   535
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P))"  (is "?t2")
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   536
proof -
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   537
  obtain n::coname where "n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   538
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   539
  then show ?t1
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   540
    by perm_simp (force simp add: fresh_prod fresh_fun_simp_OrR1 calc_atm)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   541
  obtain n::coname where "n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   542
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   543
  then show ?t2
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   544
    by perm_simp
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   545
      (force simp add: fresh_prod fresh_fun_simp_OrR1 at_prm_fresh[OF at_coname_inst] swap_simps)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   546
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   547
9be4ab2acc13 split Class.thy into parts to 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
lemma fresh_fun_simp_OrR2:
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   549
  assumes "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>b" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   550
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P) = Cut <a'>.(OrR2 <b>.M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   551
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   552
  obtain n::coname where "n\<sharp>(x,P,b,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   553
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   554
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.OrR2 <b>.M(a') (x).P, Cut <a>.OrR2 <b>.M(a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   555
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   556
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   557
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   558
    done
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   559
qed (use assms in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   560
9be4ab2acc13 split Class.thy into parts to 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
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
   562
  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
   563
  and   pi2::"coname prm"
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   564
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P) =
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   565
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(OrR2 <b>.M  a') (x).P))" (is "?t1")
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   566
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P) =
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   567
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P))"  (is "?t2")
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   568
proof -
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   569
  obtain n::coname where "n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   570
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   571
  then show ?t1
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   572
    by perm_simp (force simp add: fresh_prod fresh_fun_simp_OrR2 calc_atm)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   573
  obtain n::coname where "n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   574
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   575
  then show ?t2
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   576
    by perm_simp
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   577
      (force simp add: fresh_prod fresh_fun_simp_OrR2 at_prm_fresh[OF at_coname_inst] swap_simps)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   578
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   579
9be4ab2acc13 split Class.thy into parts to 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
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
   581
  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
   582
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P) = Cut <a'>.(ImpR (y).<b>.M a') (x).P"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   583
proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   584
  obtain n::coname where "n\<sharp>(x,P,y,b,M)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   585
    by (meson exists_fresh(2) fs_coname1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   586
  then show "\<exists>a. a \<sharp> (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P, Cut <a>.(ImpR (y).<b>.M a) (x).P)"
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   587
    apply(intro exI [where x="n"])
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   588
    apply(simp add: fresh_prod abs_fresh)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   589
    apply(fresh_guess)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   590
    done
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   591
qed (use a in \<open>fresh_guess|finite_guess\<close>)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   592
9be4ab2acc13 split Class.thy into parts to 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
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
   594
  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
   595
  and   pi2::"coname prm"
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   596
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P) =
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   597
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M  a') (x).P))" (is "?t1")
36277
9be4ab2acc13 split Class.thy into parts to 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
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P)=
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   599
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P))"  (is "?t2")
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   600
proof -
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   601
  obtain n::coname where "n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   602
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   603
  then show ?t1
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   604
    by perm_simp (force simp add: fresh_prod fresh_fun_simp_ImpR calc_atm)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   605
  obtain n::coname where "n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   606
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   607
  then show ?t2
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   608
    by perm_simp
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   609
      (force simp add: fresh_prod fresh_fun_simp_ImpR at_prm_fresh[OF at_coname_inst] swap_simps)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
   610
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   611
9be4ab2acc13 split Class.thy into parts to 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
nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)")
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
   613
  substn :: "trm \<Rightarrow> name   \<Rightarrow> coname \<Rightarrow> trm \<Rightarrow> trm" (\<open>_{_:=<_>._}\<close> [100,100,100,100] 100) 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   614
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
   615
  "(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
   616
| "\<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
   617
  (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
   618
| "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
   619
| "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
   620
  (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
   621
| "\<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
   622
  (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
   623
| "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
   624
  (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
   625
   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
   626
| "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
   627
  (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
   628
   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
   629
| "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
   630
| "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
   631
| "\<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
   632
  (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
   633
   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
   634
| "\<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
   635
| "\<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
   636
  (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
   637
   else ImpL <a>.(M{y:=<c>.P}) (x).(N{y:=<c>.P}) z)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   638
apply(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   639
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   640
apply(force simp add: fresh_prod fresh_fun_simp_NotL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   641
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   642
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   643
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   644
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   645
apply(force simp add: fresh_prod fresh_fun_simp_AndL1 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   646
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   647
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   648
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   649
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   650
apply(force simp add: fresh_prod fresh_fun_simp_AndL2 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   651
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   652
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   653
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   654
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   655
apply(force simp add: fresh_prod fresh_fun_simp_OrL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   656
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   657
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   658
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   659
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   660
apply(force simp add: fresh_prod fresh_fun_simp_OrL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   661
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   662
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   663
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   664
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   665
apply(force simp add: fresh_prod fresh_fun_simp_ImpL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   666
apply (meson exists_fresh(1) fs_name1)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   667
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   668
apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   669
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   670
apply(force simp add: fresh_prod fresh_fun_simp_ImpL abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   671
apply (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   672
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
   673
done
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
nominal_primrec (freshness_context: "(d::name,z::coname,P::trm)")
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
   676
  substc :: "trm \<Rightarrow> coname \<Rightarrow> name   \<Rightarrow> trm \<Rightarrow> trm" (\<open>_{_:=('(_'))._}\<close> [100,0,0,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   677
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
   678
  "(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
   679
| "\<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
   680
  (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
   681
| "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
   682
  (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
   683
| "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
   684
| "\<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
   685
  (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
   686
   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
   687
| "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
   688
| "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
   689
| "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
   690
  (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
   691
| "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
   692
  (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
   693
| "\<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
   694
  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
   695
| "\<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
   696
  (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
   697
   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
   698
| "\<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
   699
  ImpL <a>.(M{d:=(z).P}) (x).(N{d:=(z).P}) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   700
apply(finite_guess | simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   701
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   702
apply(force simp add: fresh_prod fresh_fun_simp_NotR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   703
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   704
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   705
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   706
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   707
apply(force simp add: fresh_prod fresh_fun_simp_AndR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   708
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   709
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   710
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   711
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   712
apply(force simp add: fresh_prod fresh_fun_simp_AndR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   713
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   714
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   715
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   716
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   717
apply(force simp add: fresh_prod fresh_fun_simp_OrR1 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   718
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   719
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   720
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   721
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   722
apply(force simp add: fresh_prod fresh_fun_simp_OrR2 abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   723
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   724
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   725
apply(simp add: abs_fresh abs_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   726
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   727
apply(force simp add: fresh_prod fresh_fun_simp_ImpR abs_fresh fresh_atm abs_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   728
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   729
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   730
apply(simp add: abs_fresh abs_supp | rule strip)+
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   731
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)")
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   732
apply(force simp add: fresh_prod fresh_fun_simp_ImpR abs_fresh fresh_atm)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   733
apply(meson exists_fresh' fin_supp)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   734
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   735
apply(simp add: abs_fresh | fresh_guess add: abs_fresh)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   736
done
9be4ab2acc13 split Class.thy into parts to 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
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
   738
36277
9be4ab2acc13 split Class.thy into parts to 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
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
   740
  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
   741
  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
   742
  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
   743
  and   "pi2\<bullet>(M{c:=(x).N}) = (pi2\<bullet>M){(pi2\<bullet>c):=(pi2\<bullet>x).(pi2\<bullet>N)}"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   744
by (nominal_induct M avoiding: c x N rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   745
   (auto simp: eq_bij fresh_bij eqvts; perm_simp)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   746
9be4ab2acc13 split Class.thy into parts to 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
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
   748
  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
   749
  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
   750
  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
   751
  and   "pi2\<bullet>(M{x:=<c>.N}) = (pi2\<bullet>M){(pi2\<bullet>x):=<(pi2\<bullet>c)>.(pi2\<bullet>N)}"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   752
by (nominal_induct M avoiding: c x N rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   753
   (auto simp: eq_bij fresh_bij eqvts; perm_simp)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   754
9be4ab2acc13 split Class.thy into parts to 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
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
   756
  shows "supp (M{y:=<c>.P}) \<subseteq> ((supp M) - {y}) \<union> (supp P)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   757
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   758
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   759
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   760
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   761
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   762
   show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   763
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   764
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   765
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   766
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   767
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   768
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   769
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   770
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   771
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   772
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   773
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   774
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   775
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   776
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   777
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   778
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   779
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   780
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   781
    by (auto simp: fs_name1 fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   782
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   783
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   784
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   785
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   786
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   787
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   788
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   789
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   790
text \<open>Identical to the previous proof\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   791
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
   792
  shows "supp (M{y:=<c>.P}) \<subseteq> supp (M) \<union> ((supp P) - {c})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   793
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   794
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   795
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   796
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   797
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   798
   show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   799
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   800
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   801
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   802
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   803
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   804
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   805
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   806
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   807
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   808
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   809
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   810
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   811
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2 fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   812
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   813
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   814
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   815
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   816
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   817
    by (auto simp: fs_name1 fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   818
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   819
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   820
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   821
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   822
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   823
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL fresh_atm; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   824
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   825
9be4ab2acc13 split Class.thy into parts to 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
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
   827
  shows "supp (M{c:=(x).P}) \<subseteq> ((supp M) - {c}) \<union> (supp P)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   828
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   829
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   830
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   831
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   832
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   833
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   834
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   835
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   836
  obtain x'::coname where x': "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   837
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   838
  with AndR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   839
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   840
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   841
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   842
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   843
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   844
  with OrR1 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   845
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   846
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   847
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   848
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   849
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   850
  with OrR2 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   851
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2; blast)  
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   852
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   853
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   854
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   855
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   856
  with ImpR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   857
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   858
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   859
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
   862
  shows "supp (M{c:=(x).P}) \<subseteq> (supp M) \<union> ((supp P) - {x})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   863
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   864
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   865
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   866
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   867
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   868
     by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   869
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   870
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   871
  obtain x'::coname where x': "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   872
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   873
  with AndR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   874
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   875
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   876
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   877
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   878
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   879
  with OrR1 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   880
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   881
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   882
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   883
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   884
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   885
  with OrR2 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   886
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2; blast)  
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   887
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   888
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   889
  obtain x'::coname where x': "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   890
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   891
  with ImpR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   892
    by (auto simp: fresh_prod abs_supp supp_atm trm.supp fs_name1 fresh_fun_simp_ImpR; blast)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   893
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   894
9be4ab2acc13 split Class.thy into parts to 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
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
   896
  shows "(supp M - {y}) \<subseteq> supp (M{y:=<c>.P})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   897
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   898
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   899
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   900
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   901
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   902
  show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   903
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   904
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   905
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   906
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   907
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   908
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   909
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   910
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   911
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   912
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   913
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   914
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   915
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   916
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   917
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   918
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   919
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   920
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   921
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   922
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   923
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   924
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   925
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   926
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   927
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   928
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   929
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   930
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   931
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   932
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   933
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   934
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   935
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   936
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   937
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   938
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   939
9be4ab2acc13 split Class.thy into parts to 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
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
   941
  shows "(supp M) \<subseteq> ((supp (M{y:=<c>.P}))::coname set)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   942
proof (nominal_induct M avoiding: y P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   943
  case (NotL coname trm name)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   944
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   945
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   946
  with NotL
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   947
  show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   948
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   949
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   950
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   951
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   952
  case (AndL1 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   953
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   954
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   955
  with AndL1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   956
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   957
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   958
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   959
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   960
  case (AndL2 name1 trm name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   961
  obtain x'::name where "x'\<sharp>(trm{y:=<c>.P},P,name1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   962
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   963
  with AndL2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   964
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   965
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   966
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   967
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   968
  case (OrL name1 trm1 name2 trm2 name3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   969
  obtain x'::name where "x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   970
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   971
  with OrL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   972
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   973
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   974
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   975
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   976
  case (ImpL coname trm1 name1 trm2 name2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   977
  obtain x'::name where "x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   978
    by (meson exists_fresh(1) fs_name1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   979
  with ImpL show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   980
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   981
           apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   982
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   983
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
   984
9be4ab2acc13 split Class.thy into parts to 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
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
   986
  shows "(supp M - {c}) \<subseteq>  supp (M{c:=(x).P})"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   987
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   988
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   989
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   990
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   991
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   992
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   993
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   994
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   995
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   996
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   997
  obtain x'::coname where "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   998
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
   999
  with AndR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1000
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1001
     apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1002
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1003
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1004
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1005
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1006
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1007
  with OrR1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1008
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1009
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1010
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1011
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1012
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1013
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1014
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1015
  with OrR2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1016
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1017
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1018
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1019
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1020
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1021
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1022
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1023
  with ImpR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1024
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1025
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1026
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1027
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1028
  
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1030
  shows "(supp M) \<subseteq> ((supp (M{c:=(x).P}))::name set)"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1031
proof (nominal_induct M avoiding: x P c rule: trm.strong_induct)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1032
  case (NotR name trm coname)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1033
  obtain x'::coname where "x'\<sharp>(trm{coname:=(x).P},P)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1034
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1035
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1036
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1037
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1038
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1039
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1040
  case (AndR coname1 trm1 coname2 trm2 coname3)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1041
  obtain x'::coname where "x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1042
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1043
  with AndR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1044
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1045
     apply (fastforce simp: fresh_def)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1046
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1047
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1048
  case (OrR1 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1049
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1050
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1051
  with OrR1 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1052
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1053
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1054
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1055
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1056
  case (OrR2 coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1057
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1058
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1059
  with OrR2 show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1060
    apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1061
     apply (auto simp: fresh_def)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1062
    done
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1063
next
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1064
  case (ImpR name coname1 trm coname2)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1065
  obtain x'::coname where "x'\<sharp>(trm{coname2:=(x).P},P,coname1)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1066
    by (meson exists_fresh'(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1067
  with ImpR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1068
    by (force simp: fresh_prod abs_supp fs_name1 supp_atm trm.supp fresh_fun_simp_ImpR)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1069
qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1070
  
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  1072
                    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
  1073
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
  1074
  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
  1075
  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
  1076
  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
  1077
  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
  1078
  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
  1079
  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
  1080
  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
  1081
  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
  1082
  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
  1083
  and   "b\<sharp>(M,P) \<Longrightarrow> b\<sharp>M{y:=<c>.P}"
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1084
  using subst_supp
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1085
  by(fastforce simp add: fresh_def supp_prod)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1086
9be4ab2acc13 split Class.thy into parts to 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
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
  1088
  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
  1089
  and   "c\<sharp>M \<Longrightarrow> M{c:=(x).P} = M"
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  1090
  by (nominal_induct M avoiding: x c P rule: trm.strong_induct)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  1091
     (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1092
9be4ab2acc13 split Class.thy into parts to 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
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
  1094
  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
  1095
  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
  1096
using a
9be4ab2acc13 split Class.thy into parts to 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
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
  1098
  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
  1099
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1100
    apply(auto simp: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1101
     apply (metis (no_types, lifting))+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1102
    done
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1103
next 
9be4ab2acc13 split Class.thy into parts to 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
  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
  1105
  then show ?case
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1106
    by (auto simp: calc_atm alpha fresh_atm abs_fresh fresh_prod fresh_left)
56073
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 53015
diff changeset
  1107
       metis
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1108
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  1110
  then show ?case
56073
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 53015
diff changeset
  1111
    by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 53015
diff changeset
  1112
      (metis crename.simps(1) crename_id crename_rename)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1113
qed (auto simp: calc_atm alpha fresh_atm abs_fresh fresh_prod fresh_left trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1114
9be4ab2acc13 split Class.thy into parts to 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
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
  1116
  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
  1117
  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
  1118
using a
9be4ab2acc13 split Class.thy into parts to 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
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
  1120
  case (NotR y M d)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1121
  obtain a::coname where "a\<sharp>(N,M{d:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1122
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1123
  with NotR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1124
    apply(auto simp: calc_atm alpha fresh_atm fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1125
    by (metis (no_types, opaque_lifting) alpha(1) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1126
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR c1 M c2 M' c3)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1128
  obtain a'::coname where "a'\<sharp>(N,M{c3:=(y).([(y,x)]\<bullet>N)},M'{c3:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,c1,c2,c3)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1129
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1130
  with AndR show ?case
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1131
    apply(auto simp: calc_atm alpha fresh_atm fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1132
    by (metis (no_types, opaque_lifting) alpha(1) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1133
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR1 d M e)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1135
  obtain a'::coname where "a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1136
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1137
  with OrR1 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1138
    by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_OrR1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1139
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR2 d M e)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1141
  obtain a'::coname where "a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1142
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1143
  with OrR2 show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1144
    by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_OrR2)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1145
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpR y d M e)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1147
  obtain a'::coname where "a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)"
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1148
    by (meson exists_fresh(2) fs_coname1)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1149
  with ImpR show ?case 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1150
    by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_ImpR) 
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1151
qed (auto simp: calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1152
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1153
lemma 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
  1154
  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
  1155
  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
  1156
using a
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1157
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
  1158
  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
  1159
  then show ?case
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1160
    apply(auto simp add: calc_atm fresh_atm abs_fresh fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1161
    by (metis (mono_tags))+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1162
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  1164
  then show ?case
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1165
    apply(auto simp add: calc_atm fresh_atm abs_fresh fresh_prod fresh_left)
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1166
    by (metis (mono_tags))+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1167
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  1169
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1170
    apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
80139
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1171
    by (metis nrename.simps(1) nrename_id nrename_rename)+
fec5a23017b5 Tidying up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 80138
diff changeset
  1172
qed (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_left abs_supp fin_supp fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1173
9be4ab2acc13 split Class.thy into parts to 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
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
  1175
  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
  1176
  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
  1177
using a
9be4ab2acc13 split Class.thy into parts to 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
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
  1179
  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
  1180
  then show ?case 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1181
    by (auto simp: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1182
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1183
  case 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
  1184
  then show ?case 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1185
    by (auto simp: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1186
next
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1187
  case (NotL d M y)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1188
  then obtain a'::name where "a'\<sharp>(N, M{x:=<c>.([(c,a)]\<bullet>N)}, [(c,a)]\<bullet>N)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1189
    by (meson exists_fresh(1) fs_name1)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1190
  with NotL show ?case
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1191
    apply(auto simp: calc_atm trm.inject fresh_atm fresh_prod fresh_left)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1192
    by (metis (no_types, opaque_lifting) alpha(2) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1193
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL x1 M x2 M' x3)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1195
  then obtain a'::name where "a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},M'{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,x1,x2,x3)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1196
    by (meson exists_fresh(1) fs_name1)  
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1197
  with OrL show ?case
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1198
    apply(auto simp: calc_atm trm.inject fresh_atm fresh_prod fresh_left)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1199
    by (metis (no_types) alpha'(2) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1200
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 u M v)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1202
  then obtain a'::name where "a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,u,v)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1203
    by (meson exists_fresh(1) fs_name1)  
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1204
  with AndL1 show ?case 
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1205
    apply(auto simp: calc_atm trm.inject fresh_atm fresh_prod fresh_left)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1206
    by (metis (mono_tags, opaque_lifting) alpha'(2) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1207
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL2 u M v)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1209
  then obtain a'::name where "a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,u,v)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1210
    by (meson exists_fresh(1) fs_name1)  
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1211
  with AndL2 show ?case 
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1212
    apply(auto simp: calc_atm trm.inject fresh_atm fresh_prod fresh_left)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1213
    by (metis (mono_tags, opaque_lifting) alpha'(2) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1214
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL d M y M' u)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1216
  then obtain a'::name where "a'\<sharp>(N,M{u:=<c>.([(c,a)]\<bullet>N)},M'{u:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,y,u)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1217
    by (meson exists_fresh(1) fs_name1)  
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1218
  with ImpL show ?case
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1219
    apply(auto simp: calc_atm trm.inject fresh_atm fresh_prod fresh_left)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1220
    by (metis (no_types) alpha'(2) trm.inject(2))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1221
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  1223
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1224
    by (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left perm_swap)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1225
qed (auto simp: calc_atm fresh_atm fresh_left)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1226
9be4ab2acc13 split Class.thy into parts to 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
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
  1228
  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
  1229
  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
  1230
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
  1231
  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
  1232
  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
  1233
  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
  1234
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
  1235
9be4ab2acc13 split Class.thy into parts to 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
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
  1237
  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
  1238
  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
  1239
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
  1240
  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
  1241
  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
  1242
  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
  1243
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
  1244
9be4ab2acc13 split Class.thy into parts to 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
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
  1246
9be4ab2acc13 split Class.thy into parts to 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
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
  1248
  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
  1249
  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
  1250
  (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
  1251
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
  1252
  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
  1253
  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
  1254
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1255
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1256
  have eq2: "(M=Ax y a) = (([(a',a)]\<bullet>M)=Ax y a')"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1257
    by (metis perm_swap(4) swap_simps(4) swap_simps(8) trm.perm(13))
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  1259
    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
  1260
  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
  1261
                              else Cut <a'>.(([(a',a)]\<bullet>M){y:=<c>.P}) (x').(([(x',x)]\<bullet>N){y:=<c>.P}))" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1262
    using fs1 fs2 by (auto simp: fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1263
  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
  1264
    using fs1 fs2 a
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1265
    unfolding eq2[symmetric]
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1266
    apply(auto simp: trm.inject)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1267
    apply(simp_all add: alpha fresh_atm fresh_prod subst_fresh eqvts perm_fresh_fresh calc_atm)
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1268
    apply (simp add: fresh_atm(2) substn_rename4)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1269
    by (metis abs_fresh(2) fresh_atm(2) fresh_prod perm_fresh_fresh(2) substn_rename4)
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  1271
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
  1272
    
9be4ab2acc13 split Class.thy into parts to 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
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
  1274
  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
  1275
  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
  1276
  (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
  1277
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
  1278
  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
  1279
  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
  1280
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1281
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1282
  have eq2: "(N=Ax x c) = (([(x',x)]\<bullet>N)=Ax x' c)"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1283
    by (metis perm_dj(1) perm_swap(1) swap_simps(1) trm.perm(1))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1284
  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
  1285
    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
  1286
  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
  1287
                              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
  1288
    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
  1289
  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
  1290
    using fs1 fs2 a
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1291
    unfolding eq2[symmetric]
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1292
    apply(auto simp: trm.inject)
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1293
    apply(simp_all add: alpha fresh_atm fresh_prod subst_fresh eqvts perm_fresh_fresh calc_atm)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1294
    by (metis abs_fresh(1) fresh_atm(1) fresh_prod perm_fresh_fresh(1) substc_rename2)
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  1296
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
  1297
9be4ab2acc13 split Class.thy into parts to 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
lemma better_Cut_substn':
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1299
  assumes "a\<sharp>[c].P" "y\<sharp>(N,x)" "M\<noteq>Ax y a"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1300
  shows "(Cut <a>.M (x).N){y:=<c>.P} = Cut <a>.(M{y:=<c>.P}) (x).N"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1301
proof -
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1302
  obtain d::name where d: "d \<sharp> (M, N, P, a, c, x, y)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1303
    by (meson exists_fresh(1) fs_name1)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1304
  then have *: "y\<sharp>([(d,x)]\<bullet>N)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1305
    by (metis assms(2) fresh_atm(1) fresh_list_cons fresh_list_nil fresh_perm_name fresh_prod)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1306
  with d have "Cut <a>.M (x).N = Cut <a>.M (d).([(d,x)]\<bullet>N)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1307
    by (metis (no_types, lifting) alpha(1) fresh_prodD perm_fresh_fresh(1) trm.inject(2))
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1308
  with * d assms show ?thesis
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1309
    apply(simp add: fresh_prod)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1310
    by (smt (verit, ccfv_SIG) forget(1) trm.inject(2))
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1311
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1312
9be4ab2acc13 split Class.thy into parts to 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
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
  1314
  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
  1315
  shows "(NotR (x).M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).M a' (z).P)"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1316
proof -
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1317
  obtain c::name where c: "c \<sharp> (M, P, d, x, z)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1318
    by (meson exists_fresh(1) fs_name1)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1319
  obtain e::coname where e: "e \<sharp> (M, P, d, x, z, c)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1320
    by (meson exists_fresh'(2) fs_coname1)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1321
  with c have "NotR (x).M d = NotR (c).([(c,x)]\<bullet>M) d"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1322
    by (metis alpha'(1) fresh_prodD(1) nrename_id nrename_swap trm.inject(3))
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1323
  with c e assms show ?thesis
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1324
    apply(simp add: fresh_prod)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1325
    by (metis forget(2) fresh_perm_app(3) trm.inject(3))
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1326
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1327
9be4ab2acc13 split Class.thy into parts to 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
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
  1329
  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
  1330
  shows "(NotL <a>.M y){y:=<c>.P} = fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1331
proof (generate_fresh "coname")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1332
  fix ca :: coname
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1333
  assume d: "ca \<sharp> (M, P, a, c, y)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1334
  then have "NotL <a>.M y = NotL <ca>.([(ca,a)]\<bullet>M) y"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1335
    by (metis alpha(2) fresh_prod perm_fresh_fresh(2) trm.inject(4))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1336
  with a d show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1337
    apply(simp add: fresh_left calc_atm forget)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1338
    apply (metis trm.inject(4))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1339
    done
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1340
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1341
9be4ab2acc13 split Class.thy into parts to 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
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
  1343
  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
  1344
  shows "(AndL1 (x).M y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M z')"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1345
proof (generate_fresh "name")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1346
  fix d:: name
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1347
  assume d: "d \<sharp> (M, P, c, x, y)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1348
  then have \<section>: "AndL1 (x).M y = AndL1 (d).([(d,x)]\<bullet>M) y"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1349
    by (metis alpha(1) fresh_prodD(1) perm_fresh_fresh(1) trm.inject(6))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1350
  with d have "(\<lambda>z'. Cut <c>.P (z').AndL1 (d).([(d, x)] \<bullet> M){x:=<c>.P} (z')) 
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1351
             = (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M (z'))"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1352
    by (metis forget(1) fresh_bij(1) fresh_prodD(1) swap_simps(1) trm.inject(6))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1353
  with d have "(\<lambda>z'. Cut <c>.P (z').AndL1 d.([(d, x)] \<bullet> M){y:=<c>.P} z')
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1354
             = (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M (z'))"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1355
    apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1356
    by (metis abs_fresh(1) assms forget(1) fresh_atm(1) fresh_aux(1) nrename_nfresh nrename_swap)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1357
  with d \<section> show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1358
    by (simp add: fresh_left calc_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1359
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1360
9be4ab2acc13 split Class.thy into parts to 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
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
  1362
  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
  1363
  shows "(AndL2 (x).M y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z')"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1364
proof (generate_fresh "name")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1365
  fix d:: name
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1366
  assume d: "d \<sharp> (M, P, c, x, y)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1367
  then have \<section>: "AndL2 (x).M y = AndL2 (d).([(d,x)]\<bullet>M) y"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1368
    by (metis alpha(1) fresh_prodD(1) perm_fresh_fresh(1) trm.inject(7))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1369
  with d have "(\<lambda>z'. Cut <c>.P (z').AndL2 (d).([(d, x)] \<bullet> M){x:=<c>.P} (z')) 
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1370
             = (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M (z'))"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1371
    by (metis forget(1) fresh_bij(1) fresh_prodD(1) swap_simps(1) trm.inject(7))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1372
  with d have "(\<lambda>z'. Cut <c>.P (z').AndL2 d.([(d, x)] \<bullet> M){y:=<c>.P} z')
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1373
             = (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M (z'))"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1374
    apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1375
    by (metis abs_fresh(1) assms forget(1) fresh_atm(1) fresh_aux(1) nrename_nfresh nrename_swap)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1376
  with d \<section> show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1377
    by (simp add: fresh_left calc_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1378
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1379
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndR_substc:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1381
  assumes "c\<sharp>([a].M,[b].N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1382
  shows "(AndR <a>.M <b>.N c){c:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.M <b>.N a') (z).P)"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1383
proof (generate_fresh "coname" , generate_fresh "coname")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1384
  fix d :: coname
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1385
    and e :: coname
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1386
  assume d: "d \<sharp> (M, N, P, a, b, c, z)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1387
    and e: "e \<sharp> (M, N, P, a, b, c, z, d)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1388
  then have "AndR <a>.M <b>.N c = AndR <d>.([(d,a)]\<bullet>M) <e>.([(e,b)]\<bullet>N) c"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1389
    by (perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) auto
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1390
  with assms d e show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1391
    apply (auto simp: fresh_left calc_atm fresh_prod fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1392
    by (metis (no_types, opaque_lifting) abs_fresh(2) forget(2) trm.inject(5))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1393
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1394
9be4ab2acc13 split Class.thy into parts to 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
lemma better_OrL_substn:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1396
  assumes "x\<sharp>([y].M,[z].N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1397
  shows "(OrL (y).M (z).N x){x:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (y).M (z).N z')"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1398
proof (generate_fresh "name" , generate_fresh "name")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1399
  fix d :: name
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1400
    and e :: name
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1401
  assume d: "d \<sharp> (M, N, P, c, x, y, z)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1402
    and e: "e \<sharp> (M, N, P, c, x, y, z, d)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1403
  with assms have "OrL (y).M (z).N x = OrL (d).([(d,y)]\<bullet>M) (e).([(e,z)]\<bullet>N) x"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1404
    by (perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) auto
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1405
  with assms d e show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1406
    apply (auto simp: fresh_left calc_atm fresh_prod fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1407
    by (metis (no_types, lifting) abs_fresh(1) forget(1) trm.inject(10))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1408
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1409
9be4ab2acc13 split Class.thy into parts to 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
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
  1411
  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
  1412
  shows "(OrR1 <a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.M a' (z).P)"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1413
proof (generate_fresh "coname")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1414
  fix c :: coname
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1415
  assume c: "c \<sharp> (M, P, a, d, z)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1416
  then have "OrR1 <a>.M d = OrR1 <c>.([(c,a)]\<bullet>M) d"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1417
    by (perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) auto
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1418
  with assms c show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1419
    apply (auto simp: fresh_left calc_atm fresh_prod fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1420
    by (metis abs_fresh(2) forget(2) trm.inject(8))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1421
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1422
9be4ab2acc13 split Class.thy into parts to 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
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
  1424
  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
  1425
  shows "(OrR2 <a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.M a' (z).P)"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1426
proof (generate_fresh "coname")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1427
  fix c :: coname
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1428
  assume c: "c \<sharp> (M, P, a, d, z)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1429
  then have "OrR2 <a>.M d = OrR2 <c>.([(c,a)]\<bullet>M) d"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1430
    by (perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) auto
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1431
  with assms c show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1432
    apply (auto simp: fresh_left calc_atm fresh_prod fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1433
    by (metis abs_fresh(2) forget(2) trm.inject(9))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1434
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1435
9be4ab2acc13 split Class.thy into parts to 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
lemma better_ImpR_substc:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1437
  assumes "d\<sharp>[a].M"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1438
  shows "(ImpR (x).<a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.M a' (z).P)"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1439
proof (generate_fresh "coname" , generate_fresh "name")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1440
  fix c :: coname
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1441
    and c' :: name
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1442
  assume c: "c \<sharp> (M, P, a, d, x, z)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1443
    and c': "c' \<sharp> (M, P, a, d, x, z, c)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1444
  have \<dagger>: "ImpR (x).<a>.M d = ImpR (c').<c>.([(c,a)]\<bullet>[(c',x)]\<bullet>M) d"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1445
    apply (rule sym)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1446
    using c c'
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1447
    apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_perm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1448
    done
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1449
  with assms c c' have "fresh_fun
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1450
        (\<lambda>a'. Cut <a'>.ImpR (c').<c>.(([(c, a)] \<bullet> [(c', x)] \<bullet> M)){d:=(z).P} a' (z).P)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1451
      = fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.M a' (z).P)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1452
    apply(intro arg_cong [where f="fresh_fun"])
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1453
    by(fastforce simp add: trm.inject alpha fresh_prod fresh_atm abs_perm fresh_left calc_atm abs_fresh forget)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1454
  with assms c c' \<dagger> show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1455
    by (auto simp: fresh_left calc_atm forget abs_fresh)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1456
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1457
9be4ab2acc13 split Class.thy into parts to 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
lemma better_ImpL_substn:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1459
  assumes "y\<sharp>(M,[x].N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1460
  shows "(ImpL <a>.M (x).N y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z')"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1461
proof (generate_fresh "coname" , generate_fresh "name")
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1462
  fix ca :: coname
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1463
    and caa :: name
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1464
  assume d: "ca \<sharp> (M, N, P, a, c, x, y)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1465
    and e: "caa \<sharp> (M, N, P, a, c, x, y, ca)"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1466
  have "ImpL <a>.M (x).N y = ImpL <ca>.([(ca,a)]\<bullet>M) (caa).([(caa,x)]\<bullet>N) y"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1467
    apply(rule sym)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1468
    using d e
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1469
    by(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_perm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1470
  with d e assms show ?thesis
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1471
    apply(simp add: fresh_left calc_atm forget abs_fresh)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1472
    apply(intro arg_cong [where f="fresh_fun"] ext)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1473
    apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm abs_fresh fresh_left calc_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1474
    by (metis forget(1) fresh_aux(1) fresh_bij(1) swap_simps(1))
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1475
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1476
9be4ab2acc13 split Class.thy into parts to 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
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
  1478
  fixes x::"name"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1479
  assumes "x\<sharp>M{c:=(y).P}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1480
  shows "x\<sharp>M"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1481
  by (meson assms fresh_def subsetD supp_subst8)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1482
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1483
lemma 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
  1484
  fixes x::"name"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1485
  assumes "x\<sharp>M{y:=<c>.P}" "x\<noteq>y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1486
  shows "x\<sharp>M"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1487
  by (meson DiffI assms fresh_def singleton_iff subset_eq supp_subst5)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  1490
  fixes a::"coname"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1491
  assumes "a\<sharp>M{c:=(y).P}" "a\<noteq>c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1492
  shows "a\<sharp>M"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1493
  by (meson Diff_iff assms fresh_def in_mono singleton_iff supp_subst7)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  1496
  fixes a::"coname"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1497
  assumes "a\<sharp>M{y:=<c>.P}" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1498
  shows "a\<sharp>M"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1499
  by (meson assms fresh_def subset_iff supp_subst6)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_crename_comm:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1502
  assumes "c\<noteq>a" "c\<noteq>b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1503
  shows "M{x:=<c>.P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{x:=<c>.(P[a\<turnstile>c>b])}"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1504
using assms
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1505
proof (nominal_induct M avoiding: x c P a b rule: trm.strong_induct)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1506
  case (Ax name coname)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1507
  then show ?case
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1508
    by(auto simp: better_crename_Cut fresh_atm)
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1509
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1510
  case (Cut coname trm1 name trm2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1511
  then show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1512
    apply(simp add: rename_fresh better_crename_Cut fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1513
    by (meson crename_ax)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1514
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1515
  case (NotL coname trm name)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1516
  obtain x'::name where "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]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1517
    by (meson exists_fresh(1) fs_name1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1518
  with NotL show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1519
    apply (simp add: subst_fresh rename_fresh trm.inject)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1520
    by (force simp: fresh_prod fresh_fun_simp_NotL better_crename_Cut fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1521
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1522
  case (AndL1 name1 trm name2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1523
  obtain x'::name where "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]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1524
    by (meson exists_fresh(1) fs_name1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1525
  with AndL1 show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1526
    apply (simp add: subst_fresh rename_fresh trm.inject)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1527
    apply (force simp: fresh_prod fresh_fun_simp_AndL1 better_crename_Cut fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1528
    done
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1529
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1530
  case (AndL2 name1 trm name2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1531
  obtain x'::name where x': "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]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1532
    by (meson exists_fresh(1) fs_name1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1533
  with AndL2 show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1534
    apply (simp add: subst_fresh rename_fresh trm.inject)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1535
    apply (auto simp: fresh_prod fresh_fun_simp_AndL2 better_crename_Cut fresh_atm)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1536
    done
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1537
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1538
  case (OrL name1 trm1 name2 trm2 name3)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1539
  obtain x'::name where x': "x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[a\<turnstile>c>b],name1,name2,
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1540
                                  trm1[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]},trm2[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1541
    by (meson exists_fresh(1) fs_name1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1542
  with OrL show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1543
    apply (simp add: subst_fresh rename_fresh trm.inject)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1544
    apply (auto simp: fresh_atm subst_fresh fresh_prod fresh_fun_simp_OrL better_crename_Cut)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1545
    done
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1546
  next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1547
  case (ImpL coname trm1 name1 trm2 name2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1548
  obtain x'::name where x': "x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[a\<turnstile>c>b],name1,name2,
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1549
                                  trm1[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]},trm2[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1550
    by (meson exists_fresh(1) fs_name1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1551
  with ImpL show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1552
    apply (simp add: subst_fresh rename_fresh trm.inject)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1553
    apply (auto simp: fresh_atm subst_fresh fresh_prod fresh_fun_simp_ImpL better_crename_Cut)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1554
    done
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1555
qed (auto simp: subst_fresh rename_fresh)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1556
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_crename_comm:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1559
  assumes "c\<noteq>a" "c\<noteq>b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1560
  shows "M{c:=(x).P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{c:=(x).(P[a\<turnstile>c>b])}"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1561
using assms
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1562
proof (nominal_induct M avoiding: x c P a b rule: trm.strong_induct)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1563
  case (Ax name coname)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1564
  then show ?case
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1565
    by(auto simp: better_crename_Cut fresh_atm)
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1566
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1567
  case (Cut coname trm1 name trm2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1568
  then show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1569
    apply(simp add: rename_fresh  better_crename_Cut)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1570
    by (meson crename_ax)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1571
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1572
  case (NotR name trm coname)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1573
  obtain c'::coname where "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]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1574
    by (meson exists_fresh' fs_coname1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1575
  with NotR show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1576
    apply(simp add: subst_fresh rename_fresh trm.inject)
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1577
    by(auto simp: fresh_prod fresh_fun_simp_NotR better_crename_Cut fresh_atm)
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1578
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1579
  case (AndR coname1 trm1 coname2 trm2 coname3)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1580
  obtain c'::coname where "c'\<sharp>(coname1,coname2,a,b,trm1{coname3:=(x).P},trm2{coname3:=(x).P},
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1581
                   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]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1582
    by (meson exists_fresh' fs_coname1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1583
  with AndR show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1584
    apply(simp add: subst_fresh rename_fresh trm.inject)
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1585
    by (auto simp: fresh_prod fresh_fun_simp_AndR better_crename_Cut subst_fresh fresh_atm)
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1586
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1587
  case (OrR1 coname1 trm coname2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1588
  obtain c'::coname where "c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],a,b, trm[a\<turnstile>c>b]{coname2:=(x).P[a\<turnstile>c>b]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1589
    by (meson exists_fresh' fs_coname1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1590
  with OrR1 show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1591
    apply(simp add: subst_fresh rename_fresh trm.inject)
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1592
    by(auto simp: fresh_prod fresh_fun_simp_OrR1 better_crename_Cut fresh_atm)
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1593
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1594
  case (OrR2 coname1 trm coname2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1595
  obtain c'::coname where "c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],a,b, trm[a\<turnstile>c>b]{coname2:=(x).P[a\<turnstile>c>b]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1596
    by (meson exists_fresh' fs_coname1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1597
  with OrR2 show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1598
    apply(simp add: subst_fresh rename_fresh trm.inject)
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1599
    by(auto simp: fresh_prod fresh_fun_simp_OrR2 better_crename_Cut fresh_atm)
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1600
next
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1601
  case (ImpR name coname1 trm coname2)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1602
  obtain c'::coname where "c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],a,b, trm[a\<turnstile>c>b]{coname2:=(x).P[a\<turnstile>c>b]})"
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1603
    by (meson exists_fresh' fs_coname1)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1604
  with ImpR show ?case
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1605
    apply(simp add: subst_fresh rename_fresh trm.inject)
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1606
    by(auto simp: fresh_prod fresh_fun_simp_ImpR better_crename_Cut fresh_atm)
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1607
qed (auto simp: subst_fresh rename_fresh trm.inject)
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1608
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  1609
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1610
lemma substn_nrename_comm:
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1611
  assumes "x\<noteq>y" "x\<noteq>z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1612
  shows "M{x:=<c>.P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{x:=<c>.(P[y\<turnstile>n>z])}"
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1613
using assms
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1614
proof (nominal_induct M avoiding: x c P y z rule: trm.strong_induct)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1615
  case (Ax name coname)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1616
  then show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1617
    by (auto simp: better_nrename_Cut fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1618
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1619
  case (Cut coname trm1 name trm2)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1620
  then show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1621
    apply(clarsimp simp: subst_fresh rename_fresh trm.inject better_nrename_Cut)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1622
    by (meson nrename_ax)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1623
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1624
  case (NotL coname trm name)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1625
  obtain x'::name where "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]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1626
    by (meson exists_fresh' fs_name1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1627
  with NotL show ?case 
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1628
    apply(clarsimp simp: rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1629
    by (auto simp add: fresh_prod fresh_fun_simp_NotL better_nrename_Cut fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1630
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1631
  case (AndL1 name1 trm name2)  
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1632
  obtain x'::name where "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)"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1633
    by (meson exists_fresh' fs_name1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1634
  with AndL1 show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1635
    apply(clarsimp simp: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1636
    by (auto simp add: fresh_prod fresh_fun_simp_AndL1 better_nrename_Cut fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1637
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1638
  case (AndL2 name1 trm name2)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1639
  obtain x'::name where "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)"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1640
    by (meson exists_fresh' fs_name1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1641
  with AndL2 show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1642
    apply(clarsimp simp: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1643
    by (auto simp add: fresh_prod fresh_fun_simp_AndL2 better_nrename_Cut fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1644
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1645
  case (OrL name1 trm1 name2 trm2 name3)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1646
  obtain x'::name where "x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[y\<turnstile>n>z],name1,name2,y,z,
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1647
                                  trm1[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]},trm2[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1648
    by (meson exists_fresh' fs_name1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1649
  with OrL show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1650
    apply (clarsimp simp: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1651
    by (auto simp add: fresh_prod fresh_fun_simp_OrL better_nrename_Cut subst_fresh fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1652
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1653
  case (ImpL coname trm1 name1 trm2 name2)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1654
  obtain x'::name where "x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[y\<turnstile>n>z],name1,name2,y,z,
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1655
                                  trm1[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]},trm2[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1656
    by (meson exists_fresh' fs_name1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1657
  with ImpL show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1658
    apply (clarsimp simp: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1659
    by (auto simp add: fresh_prod fresh_fun_simp_ImpL better_nrename_Cut subst_fresh fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1660
qed (auto simp: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1661
 
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1662
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_nrename_comm:
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1665
  assumes "x\<noteq>y" "x\<noteq>z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1666
  shows "M{c:=(x).P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{c:=(x).(P[y\<turnstile>n>z])}"
80571
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1667
using assms
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1668
proof (nominal_induct M avoiding: x c P y z rule: trm.strong_induct)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1669
  case (Ax name coname)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1670
  then show ?case 
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1671
    by (auto simp: subst_fresh rename_fresh trm.inject better_nrename_Cut fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1672
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1673
  case (Cut coname trm1 name trm2)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1674
  then show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1675
    apply (clarsimp simp: subst_fresh rename_fresh trm.inject better_nrename_Cut fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1676
    by (metis nrename_ax)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1677
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1678
  case (NotR name trm coname)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1679
  obtain c'::coname where "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]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1680
    by (meson exists_fresh' fs_coname1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1681
  with NotR show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1682
    apply(simp add: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1683
    by (auto simp add: fresh_prod fresh_fun_simp_NotR better_nrename_Cut fresh_atm)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1684
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1685
  case (AndR coname1 trm1 coname2 trm2 coname3)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1686
  obtain c'::coname where "c'\<sharp>(coname1,coname2,y,z,trm1{coname3:=(x).P},trm2{coname3:=(x).P},
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1687
                   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]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1688
    by (meson exists_fresh' fs_coname1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1689
  with AndR show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1690
    apply(simp add: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1691
    by (auto simp add: fresh_prod fresh_fun_simp_AndR better_nrename_Cut fresh_atm subst_fresh)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1692
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1693
  case (OrR1 coname1 trm coname2)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1694
  obtain c'::coname where "c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],y,z,
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1695
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>n>z]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1696
    by (meson exists_fresh' fs_coname1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1697
  with OrR1 show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1698
    apply(simp add: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1699
    by (auto simp add: fresh_prod fresh_fun_simp_OrR1 better_nrename_Cut fresh_atm subst_fresh)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1700
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1701
  case (OrR2 coname1 trm coname2)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1702
  obtain c'::coname where "c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],y,z,
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1703
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>n>z]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1704
    by (meson exists_fresh' fs_coname1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1705
  with OrR2 show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1706
    apply(simp add: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1707
    by (auto simp add: fresh_prod fresh_fun_simp_OrR2 better_nrename_Cut fresh_atm subst_fresh)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1708
next
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1709
  case (ImpR name coname1 trm coname2)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1710
  obtain c'::coname where "c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],y,z,
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1711
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>n>z]})"
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1712
    by (meson exists_fresh' fs_coname1)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1713
  with ImpR show ?case
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1714
    apply(simp add: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1715
    by (auto simp add: fresh_prod fresh_fun_simp_ImpR better_nrename_Cut fresh_atm subst_fresh)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1716
qed (auto simp: subst_fresh rename_fresh trm.inject)
673add17a05e More streamlining
paulson <lp15@cam.ac.uk>
parents: 80569
diff changeset
  1717
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_crename_comm':
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1720
  assumes "a\<noteq>c" "a\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1721
  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
  1722
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
  1723
  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
  1724
  have eq: "M{x:=<c>.P} = M{x:=<c'>.([(c',c)]\<bullet>P)}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1725
    using fs2 substn_rename4 by force
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1726
   have eq': "M[a\<turnstile>c>b]{x:=<c>.P} = M[a\<turnstile>c>b]{x:=<c'>.([(c',c)]\<bullet>P)}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1727
    using fs2 by (simp add: substn_rename4)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1728
  have eq2: "([(c',c)]\<bullet>P)[a\<turnstile>c>b] = ([(c',c)]\<bullet>P)" using fs2 assms
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1729
    by (metis crename_fresh fresh_atm(2) fresh_aux(2) fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1730
  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
  1731
  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
  1732
    using fs2
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1733
    by (simp add: fresh_atm(2) fresh_prod substn_crename_comm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  1735
  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
  1736
  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
  1737
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
  1738
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_crename_comm':
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1740
  assumes "c\<noteq>a" "c\<noteq>b" "a\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1741
  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
  1742
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
  1743
  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
  1744
  have eq: "M{c:=(x).P} = ([(c',c)]\<bullet>M){c':=(x).P}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1745
    using fs2 by (simp add: substc_rename1)
36277
9be4ab2acc13 split Class.thy into parts to 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
   have eq': "([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P} = M[a\<turnstile>c>b]{c:=(x).P}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1747
    using fs2 by (metis crename_cfresh' fresh_prod substc_rename1)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1748
  have eq2: "([(c',c)]\<bullet>M)[a\<turnstile>c>b] = ([(c',c)]\<bullet>(M[a\<turnstile>c>b]))" using fs2 assms
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1749
    by (simp add: crename_coname_eqvt fresh_atm(2) fresh_prod swap_simps(6))
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  1751
  also have "\<dots> = ([(c',c)]\<bullet>M)[a\<turnstile>c>b]{c':=(x).P[a\<turnstile>c>b]}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1752
    using fs2 assms
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1753
    by (metis crename_fresh eq eq' eq2 substc_crename_comm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P[a\<turnstile>c>b]}" using eq2 by simp
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1755
  also have "\<dots> = ([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P}" 
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1756
    using assms by (simp add: rename_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1757
  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
  1758
  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
  1759
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
  1760
9be4ab2acc13 split Class.thy into parts to 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
lemma substn_nrename_comm':
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1762
  assumes "x\<noteq>y" "x\<noteq>z" "y\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1763
  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
  1764
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
  1765
  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
  1766
  have eq: "M{x:=<c>.P} = ([(x',x)]\<bullet>M){x':=<c>.P}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1767
    using fs2 by (simp add: substn_rename3)
36277
9be4ab2acc13 split Class.thy into parts to 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
   have eq': "([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P} = M[y\<turnstile>n>z]{x:=<c>.P}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1769
    using fs2 by (metis fresh_prod nrename_nfresh' substn_rename3)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1770
  have eq2: "([(x',x)]\<bullet>M)[y\<turnstile>n>z] = ([(x',x)]\<bullet>(M[y\<turnstile>n>z]))" 
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1771
    using fs2 by (simp add: assms fresh_atm(1) fresh_prod nrename_name_eqvt swap_simps(5)) 
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  1773
  also have "\<dots> = ([(x',x)]\<bullet>M)[y\<turnstile>n>z]{x':=<c>.P[y\<turnstile>n>z]}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1774
    using fs2 by (metis assms eq eq' eq2 nrename_fresh substn_nrename_comm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = ([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P[y\<turnstile>n>z]}" using eq2 by simp
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1776
  also have "\<dots> = ([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P}" using assms by (simp add: rename_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1777
  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
  1778
  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
  1779
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
  1780
9be4ab2acc13 split Class.thy into parts to 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
lemma substc_nrename_comm':
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1782
  assumes "x\<noteq>y" "y\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1783
  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
  1784
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
  1785
  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
  1786
  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
  1787
    using fs2
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1788
    using substc_rename2 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
  1789
   have eq': "M[y\<turnstile>n>z]{c:=(x).P} = M[y\<turnstile>n>z]{c:=(x').([(x',x)]\<bullet>P)}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1790
    using fs2 by (simp add: substc_rename2)
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1791
  have eq2: "([(x',x)]\<bullet>P)[y\<turnstile>n>z] = ([(x',x)]\<bullet>P)"
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1792
    using fs2 by (metis assms(2) fresh_atm(1) fresh_aux(1) fresh_prod nrename_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
  1793
  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
  1794
  also have "\<dots> = M[y\<turnstile>n>z]{c:=(x').(([(x',x)]\<bullet>P)[y\<turnstile>n>z])}"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1795
    using fs2 by (simp add: fresh_atm(1) fresh_prod substc_nrename_comm)
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  1797
  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
  1798
  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
  1799
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
  1800
9be4ab2acc13 split Class.thy into parts to 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
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
  1802
                    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
  1803
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
  1804
                     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
  1805
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  1806
text \<open>typing contexts\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1807
41798
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 39302
diff changeset
  1808
type_synonym ctxtn = "(name\<times>ty) list"
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 39302
diff changeset
  1809
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
  1810
9be4ab2acc13 split Class.thy into parts to 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
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
  1812
  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
  1813
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
  1814
  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
  1815
| 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
  1816
9be4ab2acc13 split Class.thy into parts to 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
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
  1818
9be4ab2acc13 split Class.thy into parts to 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
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
  1820
  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
  1821
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
  1822
  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
  1823
| 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
  1824
9be4ab2acc13 split Class.thy into parts to 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
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
  1826
9be4ab2acc13 split Class.thy into parts to 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
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
  1828
  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
  1829
  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
  1830
  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
  1831
  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
  1832
  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
  1833
proof -
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1834
  show "a\<sharp>\<Gamma>" by (induct \<Gamma>) (auto simp: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty)
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1835
next
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1836
  show "x\<sharp>\<Delta>" by (induct \<Delta>) (auto simp: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1837
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
  1838
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  1839
text \<open>cut-reductions\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1840
9be4ab2acc13 split Class.thy into parts to 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
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
  1842
9be4ab2acc13 split Class.thy into parts to 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
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
  1844
  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
  1845
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
  1846
  [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
  1847
| [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
  1848
| [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
  1849
| [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
  1850
| [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
  1851
| [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
  1852
9be4ab2acc13 split Class.thy into parts to 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
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
  1854
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_Ax_elim:
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1856
  assumes "fin (Ax x a) y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1857
  shows "x=y"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1858
  using assms fin.simps trm.inject(1) 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
  1859
9be4ab2acc13 split Class.thy into parts to 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
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
  1861
  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
  1862
  shows "x=y \<and> x\<sharp>M"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1863
  using assms 
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1864
  by (cases rule: fin.cases; simp add: trm.inject; metis abs_fresh(5))
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  1867
  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
  1868
  shows "z=y \<and> z\<sharp>[x].M"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1869
  using assms by (cases rule: fin.cases; simp add: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1870
9be4ab2acc13 split Class.thy into parts to 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
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
  1872
  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
  1873
  shows "z=y \<and> z\<sharp>[x].M"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1874
  using assms by (cases rule: fin.cases; simp add: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1875
9be4ab2acc13 split Class.thy into parts to 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
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
  1877
  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
  1878
  shows "z=u \<and> z\<sharp>[x].M \<and> z\<sharp>[y].N"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1879
  using assms by (cases rule: fin.cases; simp add: trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1880
9be4ab2acc13 split Class.thy into parts to 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
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
  1882
  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
  1883
  shows "z=y \<and> z\<sharp>M \<and> z\<sharp>[x].N"
80172
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1884
  using assms
6c62605cb3f6 A little more tidying in Nominal
paulson <lp15@cam.ac.uk>
parents: 80139
diff changeset
  1885
  by (cases rule: fin.cases; simp add: trm.inject; metis abs_fresh(5))
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  1886
 
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  1889
  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
  1890
  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
  1891
  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
  1892
  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
  1893
  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
  1894
  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
  1895
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
  1896
9be4ab2acc13 split Class.thy into parts to 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
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
  1898
                   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
  1899
9be4ab2acc13 split Class.thy into parts to 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
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
  1901
  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
  1902
by (induct rule: fin.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  1903
   (auto simp: calc_atm simp add: fresh_left abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1904
9be4ab2acc13 split Class.thy into parts to 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
lemma not_fin_subst1:
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1906
  assumes "\<not>(fin M x)" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1907
  shows "\<not>(fin (M{c:=(y).P}) x)"
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1908
using assms 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1909
proof (nominal_induct M avoiding: x c y P rule: trm.strong_induct)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1910
  case (Ax name coname)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1911
  then show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1912
    using fin_rest_elims(1) substc.simps(1) by presburger
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1913
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1914
  case (Cut coname trm1 name trm2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1915
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1916
    using fin_rest_elims(1) substc.simps(1) by simp presburger
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1917
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1918
  case (NotR name trm coname)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1919
  obtain a'::coname where "a'\<sharp>(trm{coname:=(y).P},P,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1920
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1921
  with NotR show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1922
    apply (simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1923
    using fresh_fun_simp_NotR fin_rest_elims by fastforce
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1924
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1925
  case (NotL coname trm name)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1926
  then show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1927
    using fin_NotL_elim freshn_after_substc by simp blast
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1928
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1929
  case (AndR coname1 trm1 coname2 trm2 coname3)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1930
  obtain a'::coname where "a'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1931
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1932
  with AndR show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1933
    apply (simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1934
    using fresh_fun_simp_AndR fin_rest_elims by fastforce
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1935
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1936
  case (AndL1 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1937
  then show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1938
    using fin_AndL1_elim freshn_after_substc
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1939
    by simp (metis abs_fresh(1) fin.intros(3))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1940
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1941
  case (AndL2 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1942
  then show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1943
    using fin_AndL2_elim freshn_after_substc
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1944
    by simp (metis abs_fresh(1) fin.intros(4))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1945
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1946
  case (OrR1 coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1947
  obtain a'::coname where "a'\<sharp>(trm{coname2:=(y).P},coname1,P,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1948
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1949
  with OrR1 show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1950
    apply (simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1951
    using fresh_fun_simp_OrR1 fin_rest_elims by fastforce
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1952
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1953
  case (OrR2 coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1954
  obtain a'::coname where "a'\<sharp>(trm{coname2:=(y).P},coname1,P,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1955
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1956
  with OrR2 show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1957
    apply (simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1958
    using fresh_fun_simp_OrR2 fin_rest_elims by fastforce
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1959
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1960
  case (OrL name1 trm1 name2 trm2 name3)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1961
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1962
    by simp (metis abs_fresh(1) fin.intros(5) fin_OrL_elim freshn_after_substc)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1963
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1964
  case (ImpR name coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1965
  obtain a'::coname where "a'\<sharp>(trm{coname2:=(y).P},coname1,P,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1966
    by (meson exists_fresh(2) fs_coname1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1967
  with ImpR show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1968
    apply (simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1969
    using fresh_fun_simp_ImpR fin_rest_elims by fastforce
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1970
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1971
  case (ImpL coname trm1 name1 trm2 name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1972
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1973
    by simp (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim freshn_after_substc)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1974
qed
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1975
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1976
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma not_fin_subst2:
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1979
  assumes "\<not>(fin M x)" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  1980
  shows "\<not>(fin (M{y:=<c>.P}) x)"
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1981
using assms 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1982
proof (nominal_induct M avoiding: x c y P rule: trm.strong_induct)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1983
  case (Ax name coname)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1984
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1985
    using fin_rest_elims(1) substn.simps(1) by presburger
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1986
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1987
  case (Cut coname trm1 name trm2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1988
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1989
    using fin_rest_elims(1) substc.simps(1) by simp presburger
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1990
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1991
  case (NotR name trm coname)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1992
  with fin_rest_elims show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1993
    by (fastforce simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1994
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1995
  case (NotL coname trm name)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1996
  obtain a'::name where "a'\<sharp>(trm{y:=<c>.P},P,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1997
    by (meson exists_fresh(1) fs_name1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1998
  with NotL show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  1999
    apply (clarsimp simp: fresh_prod fresh_fun_simp_NotL trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2000
    by (metis fin.intros(2) fin_NotL_elim fin_rest_elims(1) freshn_after_substn)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2001
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2002
  case (AndR coname1 trm1 coname2 trm2 coname3)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2003
  obtain a'::name where "a'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2004
    by (meson exists_fresh(1) fs_name1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2005
  with AndR fin_rest_elims show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2006
    by (fastforce simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2007
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2008
  case (AndL1 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2009
  obtain a'::name where "a'\<sharp>(trm{y:=<c>.P},P,name1,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2010
    by (meson exists_fresh(1) fs_name1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2011
  with AndL1 show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2012
    apply (clarsimp simp: fresh_prod fresh_fun_simp_AndL1 trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2013
    by (metis abs_fresh(1) fin.intros(3) fin_AndL1_elim fin_rest_elims(1) freshn_after_substn)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2014
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2015
  case (AndL2 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2016
  obtain a'::name where "a'\<sharp>(trm{y:=<c>.P},P,name1,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2017
    by (meson exists_fresh(1) fs_name1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2018
  with AndL2 show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2019
    apply (clarsimp simp: fresh_prod fresh_fun_simp_AndL2 trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2020
    by (metis abs_fresh(1) fin.intros(4) fin_AndL2_elim fin_rest_elims(1) freshn_after_substn)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2021
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2022
  case (OrR1 coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2023
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2024
    using fin_rest_elims by (fastforce simp: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2025
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2026
  case (OrR2 coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2027
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2028
    using fin_rest_elims by (fastforce simp: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2029
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2030
  case (OrL name1 trm1 name2 trm2 name3)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2031
  obtain a'::name where "a'\<sharp>(trm1{y:=<c>.P},trm2{y:=<c>.P},name1,name2,P,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2032
    by (meson exists_fresh(1) fs_name1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2033
  with OrL show ?case  
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2034
    apply (simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2035
    by (metis (full_types) abs_fresh(1) fin.intros(5) fin_OrL_elim fin_rest_elims(1) fresh_fun_simp_OrL freshn_after_substn)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2036
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2037
  case (ImpR name coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2038
  with fin_rest_elims show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2039
    by (fastforce simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2040
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2041
  case (ImpL coname trm1 name1 trm2 name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2042
  obtain a'::name where "a'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},name1,P,x)"
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2043
    by (meson exists_fresh(1) fs_name1)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2044
  with ImpL show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2045
    apply (simp add: fresh_prod trm.inject)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2046
    by (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim fin_rest_elims(1) fresh_fun_simp_ImpL freshn_after_substn)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2047
qed
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2048
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_subst1:
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2051
  assumes "fin M x" "x\<noteq>y" "x\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2052
  shows "fin (M{y:=<c>.P}) x"
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2053
  using assms
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2054
proof (nominal_induct M avoiding: x y c P rule: trm.strong_induct)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2055
  case (AndL1 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2056
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2057
    apply (clarsimp simp add: subst_fresh abs_fresh dest!: fin_AndL1_elim)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2058
    by (metis abs_fresh(1) fin.intros(3) fresh_prod subst_fresh(3))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2059
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2060
  case (AndL2 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2061
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2062
    apply (clarsimp simp add: subst_fresh abs_fresh dest!: fin_AndL2_elim)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2063
    by (metis abs_fresh(1) fin.intros(4) fresh_prod subst_fresh(3))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2064
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2065
  case (OrR1 coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2066
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2067
    by (auto simp add: subst_fresh abs_fresh dest!: fin_rest_elims)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2068
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2069
  case (OrR2 coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2070
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2071
    by (auto simp add: subst_fresh abs_fresh dest!: fin_rest_elims)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2072
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2073
  case (OrL name1 trm1 name2 trm2 name3)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2074
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2075
    apply (clarsimp simp add: subst_fresh abs_fresh)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2076
    by (metis abs_fresh(1) fin.intros(5) fin_OrL_elim fresh_prod subst_fresh(3))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2077
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2078
  case (ImpR name coname1 trm coname2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2079
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2080
    by (auto simp add: subst_fresh abs_fresh dest!: fin_rest_elims)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2081
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2082
  case (ImpL coname trm1 name1 trm2 name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2083
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2084
    apply (clarsimp simp add: subst_fresh abs_fresh)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2085
    by (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim fresh_prod subst_fresh(3))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2086
qed (auto dest!: fin_elims simp add: subst_fresh abs_fresh)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2087
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2088
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2089
thm abs_fresh fresh_prod
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2090
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_subst2:
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2092
  assumes "fin M y" "x\<noteq>y" "y\<sharp>P" "M\<noteq>Ax y c" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2093
  shows "fin (M{c:=(x).P}) y"
80593
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2094
using assms
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2095
proof (nominal_induct M avoiding: x y c P rule: trm.strong_induct)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2096
  case (Ax name coname)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2097
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2098
    using fin_Ax_elim by force
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2099
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2100
  case (NotL coname trm name)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2101
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2102
    by simp (metis fin.intros(2) fin_NotL_elim fresh_prod subst_fresh(5)) 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2103
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2104
  case (AndL1 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2105
  then show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2106
    by simp (metis abs_fresh(1) fin.intros(3) fin_AndL1_elim fresh_prod subst_fresh(5))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2107
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2108
  case (AndL2 name1 trm name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2109
  then show ?case 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2110
    by simp (metis abs_fresh(1) fin.intros(4) fin_AndL2_elim fresh_prod subst_fresh(5)) 
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2111
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2112
  case (OrL name1 trm1 name2 trm2 name3)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2113
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2114
    by simp (metis abs_fresh(1) fin.intros(5) fin_OrL_elim fresh_prod subst_fresh(5))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2115
next
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2116
  case (ImpL coname trm1 name1 trm2 name2)
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2117
  then show ?case
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2118
    by simp (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim fresh_prod subst_fresh(5))
f2eb4fa95525 more proof tidying
paulson <lp15@cam.ac.uk>
parents: 80571
diff changeset
  2119
qed (use fin_rest_elims in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2120
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2121
36277
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_substn_nrename:
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2123
  assumes "fin M x" "x\<noteq>y" "x\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2124
  shows "M[x\<turnstile>n>y]{y:=<c>.P} = Cut <c>.P (x).(M{y:=<c>.P})"
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2125
using assms [[simproc del: defined_all]]
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2126
proof (nominal_induct M avoiding: x y c P rule: trm.strong_induct)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2127
  case (Ax name coname)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2128
  then show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2129
    by (metis fin_Ax_elim fresh_atm(1,3) fresh_prod nrename_swap subst_rename(3) substn.simps(1) trm.fresh(1))
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2130
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2131
  case (NotL coname trm name)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2132
  obtain z::name where "z \<sharp> (trm,y,x,P,trm[x\<turnstile>n>y]{y:=<c>.P})"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2133
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2134
  with NotL show ?case 
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2135
    apply (clarsimp simp add: fresh_prod fresh_fun_simp_NotL trm.inject alpha fresh_atm calc_atm abs_fresh)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2136
    by (metis fin_NotL_elim nrename_fresh nrename_swap substn_nrename_comm')
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2137
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2138
  case (AndL1 name1 trm name2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2139
  obtain z::name where "z \<sharp> (name2,name1,P,trm[name2\<turnstile>n>y]{y:=<c>.P},y,P,trm)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2140
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2141
  with AndL1 show ?case 
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2142
    using fin_AndL1_elim[OF AndL1.prems(1)]
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2143
    by simp (metis abs_fresh(1) fresh_atm(1) fresh_fun_simp_AndL1 fresh_prod nrename_fresh subst_fresh(3))
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2144
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2145
  case (AndL2 name1 trm name2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2146
  obtain z::name where "z \<sharp> (name2,name1,P,trm[name2\<turnstile>n>y]{y:=<c>.P},y,P,trm)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2147
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2148
  with AndL2 show ?case 
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2149
    using fin_AndL2_elim[OF AndL2.prems(1)]
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2150
    by simp (metis abs_fresh(1) fresh_atm(1) fresh_fun_simp_AndL2 fresh_prod nrename_fresh subst_fresh(3))
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2151
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2152
  case (OrL name1 trm1 name2 trm2 name3)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2153
  obtain z::name where "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)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2154
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2155
  with OrL show ?case 
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2156
    using fin_OrL_elim[OF OrL.prems(1)]
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2157
    by (auto simp add: abs_fresh fresh_fun_simp_OrL fresh_atm(1) nrename_fresh subst_fresh(3))
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2158
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2159
  case (ImpL coname trm1 name1 trm2 name2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2160
  obtain z::name where "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)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2161
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2162
  with ImpL show ?case 
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2163
    using fin_ImpL_elim[OF ImpL.prems(1)]
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2164
    by (auto simp add: abs_fresh fresh_fun_simp_ImpL fresh_atm(1) nrename_fresh subst_fresh(3))
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2165
qed (use fin_rest_elims in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2166
9be4ab2acc13 split Class.thy into parts to 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
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
  2168
  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
  2169
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
  2170
  [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
  2171
| [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
  2172
| [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
  2173
| [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
  2174
| [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
  2175
| [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
  2176
9be4ab2acc13 split Class.thy into parts to 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
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
  2178
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_Ax_elim:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2180
  assumes "fic (Ax x a) b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2181
  shows "a=b"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2182
  using assms by (auto simp: trm.inject elim!: fic.cases)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_NotR_elim:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2185
  assumes "fic (NotR (x).M a) b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2186
  shows "a=b \<and> b\<sharp>M"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2187
  using assms
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2188
  by (auto simp: trm.inject elim!: fic.cases) (metis abs_fresh(6))
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_OrR1_elim:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2191
  assumes "fic (OrR1 <a>.M b) c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2192
  shows "b=c \<and> c\<sharp>[a].M"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2193
  using assms by (auto simp: trm.inject elim!: fic.cases)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_OrR2_elim:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2196
  assumes "fic (OrR2 <a>.M b) c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2197
  shows "b=c \<and> c\<sharp>[a].M"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2198
  using assms by (auto simp: trm.inject elim!: fic.cases)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_AndR_elim:
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2201
  assumes "fic (AndR <a>.M <b>.N c) d"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2202
  shows "c=d \<and> d\<sharp>[a].M \<and> d\<sharp>[b].N"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2203
  using assms by (auto simp: trm.inject elim!: fic.cases)
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  2206
  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
  2207
  shows "b=c \<and> b\<sharp>[a].M"
80569
f1872ef41766 Revised mixfix and streamlined proofs
paulson <lp15@cam.ac.uk>
parents: 80172
diff changeset
  2208
  using assms by (auto simp: trm.inject elim!: fic.cases) (metis abs_fresh(6))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2209
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2210
lemma 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
  2211
  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
  2212
  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
  2213
  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
  2214
  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
  2215
  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
  2216
  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
  2217
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
  2218
9be4ab2acc13 split Class.thy into parts to 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
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
  2220
                   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
  2221
9be4ab2acc13 split Class.thy into parts to 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
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
  2223
  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
  2224
by (induct rule: fic.induct)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2225
   (auto simp: calc_atm simp add: fresh_left abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2226
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2227
lemma not_fic_subst1:
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2228
  assumes "\<not>(fic M a)" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2229
  shows "\<not>(fic (M{y:=<c>.P}) a)"
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2230
using assms
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2231
proof (nominal_induct M avoiding: a c y P rule: trm.strong_induct)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2232
  case (Ax name coname)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2233
  then show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2234
    using fic_rest_elims(1) substn.simps(1) by presburger
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2235
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2236
  case (Cut coname trm1 name trm2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2237
  then show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2238
    by (metis abs_fresh(2) better_Cut_substn fic_rest_elims(1) fresh_prod)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2239
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2240
  case (NotR name trm coname)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2241
  then show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2242
    by (metis fic.intros(2) fic_NotR_elim fresh_prod freshc_after_substn substn.simps(3))
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2243
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2244
  case (NotL coname trm name)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2245
  obtain x'::name where "x' \<sharp> (trm{y:=<c>.P},P,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2246
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2247
  with NotL show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2248
    by (simp add: fic.intros fresh_prod) (metis fic_rest_elims(1,2) fresh_fun_simp_NotL)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2249
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2250
  case (AndR coname1 trm1 coname2 trm2 coname3)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2251
  then show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2252
    by simp (metis abs_fresh(2) fic.intros(3) fic_AndR_elim freshc_after_substn)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2253
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2254
  case (AndL1 name1 trm name2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2255
  obtain x'::name where "x' \<sharp> (trm{y:=<c>.P},P,name1,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2256
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2257
  with AndL1 fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2258
    by (simp add: fic.intros fresh_prod)(metis fresh_fun_simp_AndL1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2259
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2260
  case (AndL2 name1 trm name2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2261
  obtain x'::name where "x' \<sharp> (trm{y:=<c>.P},P,name1,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2262
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2263
  with AndL2 fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2264
    by (simp add: fic.intros fresh_prod) (metis fresh_fun_simp_AndL2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2265
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2266
  case (OrR1 coname1 trm coname2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2267
  then show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2268
    by simp (metis abs_fresh(2) fic.simps fic_OrR1_elim freshc_after_substn)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2269
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2270
  case (OrR2 coname1 trm coname2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2271
  then show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2272
    by simp (metis abs_fresh(2) fic.simps fic_OrR2_elim freshc_after_substn)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2273
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2274
  case (OrL name1 trm1 name2 trm2 name3)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2275
  obtain x'::name where "x' \<sharp> (trm1{y:=<c>.P},trm2{y:=<c>.P},P,name1,name2,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2276
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2277
  with OrL fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2278
    by (simp add: fic.intros fresh_prod) (metis fresh_fun_simp_OrL)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2279
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2280
  case (ImpR name coname1 trm coname2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2281
  then show ?case 
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2282
    by simp (metis abs_fresh(2) fic.intros(6) fic_ImpR_elim freshc_after_substn)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2283
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2284
  case (ImpL coname trm1 name1 trm2 name2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2285
  obtain x'::name where "x' \<sharp> (trm1{name2:=<c>.P},trm2{name2:=<c>.P},P,name1,name2,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2286
    by (meson exists_fresh(1) fs_name1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2287
  with ImpL fic_rest_elims fresh_fun_simp_ImpL show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2288
    by (simp add: fresh_prod) fastforce
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2289
qed
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2290
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2291
lemma not_fic_subst2:
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2292
  assumes "\<not>(fic M a)" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2293
  shows "\<not>(fic (M{c:=(y).P}) a)"
80595
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2294
using assms
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2295
proof (nominal_induct M avoiding: a c y P rule: trm.strong_induct)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2296
  case (NotR name trm coname)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2297
  obtain c'::coname where "c'\<sharp>(trm{coname:=(y).P},P,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2298
    by (meson exists_fresh'(2) fs_coname1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2299
  with NotR fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2300
    apply (clarsimp simp: fresh_prod fresh_fun_simp_NotR)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2301
    by (metis fic.intros(2) fic_NotR_elim freshc_after_substc)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2302
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2303
  case (AndR coname1 trm1 coname2 trm2 coname3)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2304
  obtain c'::coname where "c'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2305
    by (meson exists_fresh'(2) fs_coname1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2306
  with AndR fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2307
    apply (clarsimp simp: fresh_prod fresh_fun_simp_AndR)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2308
    by (metis abs_fresh(2) fic.intros(3) fic_AndR_elim freshc_after_substc)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2309
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2310
  case (OrR1 coname1 trm coname2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2311
  obtain c'::coname where "c'\<sharp>(trm{coname2:=(y).P},P,coname1,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2312
    by (meson exists_fresh'(2) fs_coname1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2313
  with OrR1 fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2314
    apply (clarsimp simp: fresh_prod fresh_fun_simp_OrR1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2315
    by (metis abs_fresh(2) fic.intros(4) fic_OrR1_elim freshc_after_substc)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2316
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2317
  case (OrR2 coname1 trm coname2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2318
  obtain c'::coname where "c'\<sharp>(trm{coname2:=(y).P},P,coname1,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2319
    by (meson exists_fresh'(2) fs_coname1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2320
  with OrR2 fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2321
    apply (clarsimp simp: fresh_prod fresh_fun_simp_OrR2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2322
    by (metis abs_fresh(2) fic.simps fic_OrR2_elim freshc_after_substc)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2323
next
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2324
  case (ImpR name coname1 trm coname2)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2325
  obtain c'::coname where "c'\<sharp>(trm{coname2:=(y).P},P,coname1,a)"
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2326
    by (meson exists_fresh'(2) fs_coname1)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2327
  with ImpR fic_rest_elims show ?case
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2328
    apply (clarsimp simp: fresh_prod fresh_fun_simp_ImpR)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2329
    by (metis abs_fresh(2) fic.intros(6) fic_ImpR_elim freshc_after_substc)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2330
qed (use fic_rest_elims in auto)
1effd04264cc Got rid of another 250 apply-lines
paulson <lp15@cam.ac.uk>
parents: 80593
diff changeset
  2331
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2332
lemma fic_subst1:
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2333
  assumes "fic M a" "a\<noteq>b" "a\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2334
  shows "fic (M{b:=(x).P}) a"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2335
using assms
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2336
proof (nominal_induct M avoiding: x b a P rule: trm.strong_induct)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2337
  case (Ax name coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2338
  then show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2339
    using fic_Ax_elim by force
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2340
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2341
  case (NotR name trm coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2342
  with fic_NotR_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2343
    by (fastforce simp add: fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2344
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2345
  case (AndR coname1 trm1 coname2 trm2 coname3)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2346
  with fic_AndR_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2347
    by (fastforce simp add: abs_fresh fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2348
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2349
  case (OrR1 coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2350
  with fic_OrR1_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2351
    by (fastforce simp add: abs_fresh fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2352
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2353
  case (OrR2 coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2354
  with fic_OrR2_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2355
    by (fastforce simp add: abs_fresh fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2356
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2357
  case (ImpR name coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2358
  with fic_ImpR_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2359
    by (fastforce simp add: abs_fresh fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2360
qed (use fic_rest_elims in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2361
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2362
lemma fic_subst2:
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2363
  assumes "fic M a" "c\<noteq>a" "a\<sharp>P" "M\<noteq>Ax x a" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2364
  shows "fic (M{x:=<c>.P}) a"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2365
using assms
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2366
proof (nominal_induct M avoiding: x a c P rule: trm.strong_induct)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2367
  case (Ax name coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2368
  then show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2369
    using fic_Ax_elim by force
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2370
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2371
  case (NotR name trm coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2372
  with fic_NotR_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2373
    by (fastforce simp add: fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2374
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2375
  case (AndR coname1 trm1 coname2 trm2 coname3)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2376
  with fic_AndR_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2377
    by simp (metis abs_fresh(2) crename_cfresh crename_fresh fic.intros(3) fresh_atm(2) substn_crename_comm')
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2378
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2379
  case (OrR1 coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2380
  with fic_OrR1_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2381
    by (fastforce simp add: abs_fresh fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2382
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2383
  case (OrR2 coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2384
  with fic_OrR2_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2385
    by (fastforce simp add: abs_fresh fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2386
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2387
  case (ImpR name coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2388
  with fic_ImpR_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2389
    by (fastforce simp add: abs_fresh fresh_prod subst_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2390
qed (use fic_rest_elims in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2391
9be4ab2acc13 split Class.thy into parts to 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
lemma fic_substc_crename:
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2393
  assumes "fic M a" "a\<noteq>b" "a\<sharp>P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2394
  shows "M[a\<turnstile>c>b]{b:=(y).P} = Cut <a>.(M{b:=(y).P}) (y).P"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2395
using assms
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2396
proof (nominal_induct M avoiding: a b  y P rule: trm.strong_induct)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2397
  case (Ax name coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2398
  with fic_Ax_elim show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2399
    by(force simp add: trm.inject alpha(2) fresh_atm(2,4) swap_simps(4,8))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2400
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2401
  case (Cut coname trm1 name trm2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2402
  with fic_rest_elims show ?case by force
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2403
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2404
  case (NotR name trm coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2405
  with fic_NotR_elim[OF NotR.prems(1)] show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2406
    by (simp add: trm.inject crename_fresh fresh_fun_simp_NotR subst_fresh(6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2407
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2408
  case (AndR coname1 trm1 coname2 trm2 coname3)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2409
  with AndR fic_AndR_elim[OF AndR.prems(1)] show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2410
    by (simp add: abs_fresh rename_fresh fresh_fun_simp_AndR fresh_atm(2) subst_fresh(6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2411
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2412
  case (OrR1 coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2413
  with fic_OrR1_elim[OF OrR1.prems(1)] show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2414
    by (simp add: abs_fresh rename_fresh fresh_fun_simp_OrR1 fresh_atm(2) subst_fresh(6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2415
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2416
  case (OrR2 coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2417
  with fic_OrR2_elim[OF OrR2.prems(1)] show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2418
    by (simp add: abs_fresh rename_fresh fresh_fun_simp_OrR2 fresh_atm(2) subst_fresh(6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2419
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2420
  case (ImpR name coname1 trm coname2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2421
  with fic_ImpR_elim[OF ImpR.prems(1)] crename_fresh show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2422
    by (force simp add: abs_fresh fresh_fun_simp_ImpR fresh_atm(2) subst_fresh(6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2423
qed (use fic_rest_elims in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2424
9be4ab2acc13 split Class.thy into parts to 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
inductive
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
  2426
  l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>l _\<close> [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2427
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2428
  LAxR:  "\<lbrakk>x\<sharp>M; a\<sharp>b; fic M a\<rbrakk> \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^sub>l M[a\<turnstile>c>b]"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2429
| LAxL:  "\<lbrakk>a\<sharp>M; x\<sharp>y; fin M x\<rbrakk> \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^sub>l M[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2430
| LNot:  "\<lbrakk>y\<sharp>(M,N); x\<sharp>(N,y); a\<sharp>(M,N,b); b\<sharp>M; y\<noteq>x; b\<noteq>a\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2431
          Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^sub>l Cut <b>.N (x).M" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2432
| LAnd1: "\<lbrakk>b\<sharp>([a1].M1,[a2].M2,N,a1,a2); y\<sharp>([x].N,M1,M2,x); x\<sharp>(M1,M2); a1\<sharp>(M2,N); a2\<sharp>(M1,N); a1\<noteq>a2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2433
          Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^sub>l Cut <a1>.M1 (x).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2434
| LAnd2: "\<lbrakk>b\<sharp>([a1].M1,[a2].M2,N,a1,a2); y\<sharp>([x].N,M1,M2,x); x\<sharp>(M1,M2); a1\<sharp>(M2,N); a2\<sharp>(M1,N); a1\<noteq>a2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2435
          Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL2 (x).N y) \<longrightarrow>\<^sub>l Cut <a2>.M2 (x).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2436
| LOr1:  "\<lbrakk>b\<sharp>([a].M,N1,N2,a); y\<sharp>([x1].N1,[x2].N2,M,x1,x2); x1\<sharp>(M,N2); x2\<sharp>(M,N1); a\<sharp>(N1,N2); x1\<noteq>x2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2437
          Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x1).N1"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2438
| LOr2:  "\<lbrakk>b\<sharp>([a].M,N1,N2,a); y\<sharp>([x1].N1,[x2].N2,M,x1,x2); x1\<sharp>(M,N2); x2\<sharp>(M,N1); a\<sharp>(N1,N2); x1\<noteq>x2\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2439
          Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x2).N2"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2440
| 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
  2441
          c\<sharp>(P,[a].M,b,a); a\<sharp>([c].N,P); y\<sharp>(N,[x].M)\<rbrakk> \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2442
          Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^sub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2443
9be4ab2acc13 split Class.thy into parts to 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
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
  2445
9be4ab2acc13 split Class.thy into parts to 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
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
  2447
  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
  2448
  and   pi2::"coname prm"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2449
  shows "(pi1\<bullet>M) \<longrightarrow>\<^sub>l (pi1\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^sub>l M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2450
  and   "(pi2\<bullet>M) \<longrightarrow>\<^sub>l (pi2\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^sub>l M'"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2451
  using l_redu.eqvt perm_pi_simp by metis+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2452
9be4ab2acc13 split Class.thy into parts to 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
nominal_inductive l_redu
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2454
  by (force simp add: abs_fresh fresh_atm rename_fresh fresh_prod abs_supp fin_supp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2455
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2456
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2457
lemma fresh_l_redu_x:
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2458
  fixes z::"name"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2459
  shows "M \<longrightarrow>\<^sub>l M' \<Longrightarrow> z\<sharp>M \<Longrightarrow> z\<sharp>M'"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2460
proof (induct rule: l_redu.induct)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2461
  case (LAxL a M x y)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2462
  then show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2463
    by (metis abs_fresh(1,5) nrename_nfresh nrename_rename trm.fresh(1,3))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2464
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2465
  case (LImp z N y P x M b a c)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2466
  then show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2467
    apply (simp add: abs_fresh fresh_prod)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2468
    by (metis abs_fresh(3,5) abs_supp(5) fs_name1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2469
qed (auto simp add: abs_fresh fresh_prod crename_nfresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2470
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2471
lemma fresh_l_redu_a:
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2472
  fixes c::"coname"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2473
  shows "M \<longrightarrow>\<^sub>l M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2474
proof (induct rule: l_redu.induct)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2475
  case (LAxR x M a b)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2476
  then show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2477
    apply (simp add: abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2478
    by (metis crename_cfresh crename_rename)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2479
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2480
  case (LAxL a M x y)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2481
  with nrename_cfresh show ?case
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2482
    by (simp add: abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2483
qed (auto simp add: abs_fresh fresh_prod crename_nfresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2484
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2485
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2486
lemmas fresh_l_redu = fresh_l_redu_x fresh_l_redu_a
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2487
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LAxR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2489
  shows "fic M a \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^sub>l M[a\<turnstile>c>b]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2490
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
  2491
  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
  2492
  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
  2493
  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
  2494
  have "Cut <a>.M (x).(Ax x b) =  Cut <a'>.([(a',a)]\<bullet>M) (x').(Ax x' b)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2495
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2496
  also have "\<dots> \<longrightarrow>\<^sub>l ([(a',a)]\<bullet>M)[a'\<turnstile>c>b]" using fs1 fs2 fin
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2497
    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
  2498
  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
  2499
  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
  2500
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
  2501
    
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LAxL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2503
  shows "fin M x \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^sub>l M[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2504
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
  2505
  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
  2506
  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
  2507
  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
  2508
  have "Cut <a>.(Ax y a) (x).M = Cut <a'>.(Ax y a') (x').([(x',x)]\<bullet>M)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2509
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2510
  also have "\<dots> \<longrightarrow>\<^sub>l ([(x',x)]\<bullet>M)[x'\<turnstile>n>y]" using fs1 fs2 fin
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2511
    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
  2512
  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
  2513
  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
  2514
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
  2515
9be4ab2acc13 split Class.thy into parts to 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
lemma better_LNot_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2517
  shows "\<lbrakk>y\<sharp>N; a\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^sub>l Cut <b>.N (x).M"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2518
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
  2519
  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
  2520
  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
  2521
  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
  2522
  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
  2523
  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
  2524
  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
  2525
                      = 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
  2526
    using f1 f2 f3 f4 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2527
    by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2528
  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
  2529
    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
  2530
  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
  2531
    using f1 f2 f3 f4 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2532
    by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2533
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <b'>.([(b',b)]\<bullet>N) (x').([(x',x)]\<bullet>M)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2534
    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
  2535
    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
  2536
  also have "\<dots> = Cut <b>.N (x).M"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2537
    using f1 f2 f3 f4 by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2538
  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
  2539
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
  2540
9be4ab2acc13 split Class.thy into parts to 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
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
  2542
  shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2543
         \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).N y) \<longrightarrow>\<^sub>l Cut <b1>.M1 (x).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2544
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
  2545
  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
  2546
  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
  2547
  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
  2548
  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
  2549
  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
  2550
  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
  2551
  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
  2552
                      = 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
  2553
    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
  2554
    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
  2555
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2556
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2557
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  2559
                                                               (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
  2560
    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
  2561
    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
  2562
    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
  2563
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2564
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <b1'>.([(b1',b1)]\<bullet>M1) (x').([(x',x)]\<bullet>N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2565
    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
  2566
    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
  2567
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2568
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2569
    done
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = Cut <b1>.M1 (x).N"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2571
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2572
  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
  2573
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
  2574
9be4ab2acc13 split Class.thy into parts to 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
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
  2576
  shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2577
         \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).N y) \<longrightarrow>\<^sub>l Cut <b2>.M2 (x).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2578
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
  2579
  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
  2580
  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
  2581
  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
  2582
  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
  2583
  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
  2584
  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
  2585
  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
  2586
                      = 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
  2587
    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
  2588
    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
  2589
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2590
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2591
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  2593
                                                               (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
  2594
    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
  2595
    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
  2596
    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
  2597
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2598
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <b2'>.([(b2',b2)]\<bullet>M2) (x').([(x',x)]\<bullet>N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2599
    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
  2600
    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
  2601
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2602
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2603
    done
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = Cut <b2>.M2 (x).N"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2605
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2606
  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
  2607
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
  2608
9be4ab2acc13 split Class.thy into parts to 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
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
  2610
  shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2611
         \<Longrightarrow> Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x1).N1"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2612
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
  2613
  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
  2614
  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
  2615
  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
  2616
  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
  2617
  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
  2618
  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
  2619
  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
  2620
                      = 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
  2621
    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
  2622
    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
  2623
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2624
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2625
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  2627
              (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
  2628
    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
  2629
    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
  2630
    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
  2631
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2632
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.([(a',a)]\<bullet>M) (x1').([(x1',x1)]\<bullet>N1)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2633
    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
  2634
    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
  2635
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2636
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2637
    done
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = Cut <a>.M (x1).N1"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2639
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2640
  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
  2641
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
  2642
9be4ab2acc13 split Class.thy into parts to 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
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
  2644
  shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2645
         \<Longrightarrow> Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x2).N2"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2646
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
  2647
  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
  2648
  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
  2649
  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
  2650
  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
  2651
  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
  2652
  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
  2653
  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
  2654
                      = 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
  2655
    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
  2656
    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
  2657
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2658
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2659
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  2661
              (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
  2662
    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
  2663
    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
  2664
    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
  2665
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2666
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.([(a',a)]\<bullet>M) (x2').([(x2',x2)]\<bullet>N2)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2667
    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
  2668
    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
  2669
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2670
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2671
    done
9be4ab2acc13 split Class.thy into parts to 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
  also have "\<dots> = Cut <a>.M (x2).N2"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2673
    using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2674
  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
  2675
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
  2676
9be4ab2acc13 split Class.thy into parts to 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
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
  2678
  shows "\<lbrakk>z\<sharp>(N,[y].P); b\<sharp>[a].M; a\<sharp>N\<rbrakk> 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2679
         \<Longrightarrow> Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^sub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2680
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
  2681
  assume fs: "z\<sharp>(N,[y].P)" "b\<sharp>[a].M" "a\<sharp>N"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2682
  obtain y'::"name" and x'::"name" and z'::"name" 
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2683
    where f1: "y'\<sharp>(y,M,N,P,z,x)" and f2: "x'\<sharp>(y,M,N,P,z,x,y')" and f3: "z'\<sharp>(y,M,N,P,z,x,y',x')"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2684
    by (meson exists_fresh(1) fs_name1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2685
  obtain a'::"coname" and b'::"coname" and c'::"coname" 
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2686
    where f4: "a'\<sharp>(a,N,P,M,b)" and f5: "b'\<sharp>(a,N,P,M,b,c,a')" and f6: "c'\<sharp>(a,N,P,M,b,c,a',b')"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2687
    by (meson exists_fresh(2) fs_coname1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2688
  have "Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2689
                      =  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
  2690
    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
  2691
    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
  2692
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2693
    apply(auto simp: perm_fresh_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2694
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  2696
                           (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
  2697
    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
  2698
    apply(rule_tac sym)
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2699
    apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2700
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  2701
  also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.(Cut <c'>.([(c',c)]\<bullet>N) (x').([(a',a)]\<bullet>[(x',x)]\<bullet>M)) (y').([(y',y)]\<bullet>P)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2702
    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
  2703
    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
  2704
    apply(rule l_redu.intros)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  2705
    apply(auto simp: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2706
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  2708
    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
  2709
    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
  2710
    apply(perm_simp add: alpha fresh_prod fresh_atm)
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2711
    apply(simp add: trm.inject abs_fresh alpha(1) fresh_perm_app(4) perm_compose(4) perm_dj(4))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2712
    apply (metis alpha'(2) crename_fresh crename_swap swap_simps(2,4,6))
36277
9be4ab2acc13 split Class.thy into parts to 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
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  2715
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
  2716
9be4ab2acc13 split Class.thy into parts to 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
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
  2718
  fixes M::"trm"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2719
    and   a::"coname"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2720
  assumes "[a].M = [b].N" "c\<sharp>(a,b,M,N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2721
  shows "M = [(a,c)]\<bullet>[(b,c)]\<bullet>N"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2722
  by (metis alpha_fresh'(2) assms fresh_atm(2) fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2723
9be4ab2acc13 split Class.thy into parts to 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
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
  2725
  fixes M::"trm"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2726
    and   x::"name"
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2727
  assumes "[x].M = [y].N" "z\<sharp>(x,y,M,N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2728
  shows "M = [(x,z)]\<bullet>[(y,z)]\<bullet>N"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2729
  by (metis alpha_fresh'(1) assms fresh_atm(1) fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2730
9be4ab2acc13 split Class.thy into parts to 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
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
  2732
  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
  2733
  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
  2734
  and   a::"coname"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2735
assumes "[x].[b].M = [y].[c].N" "z\<sharp>(x,y,M,N)" "a\<sharp>(b,c,M,N)"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2736
  shows "M = [(x,z)]\<bullet>[(b,a)]\<bullet>[(c,a)]\<bullet>[(y,z)]\<bullet>N"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2737
  using assms
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2738
  apply(clarsimp simp: alpha_fresh fresh_prod fresh_atm 
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2739
      abs_supp fin_supp abs_fresh abs_perm fresh_left calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2740
  by (metis perm_swap(1,2))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2741
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma Cut_l_redu_elim:
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2744
  assumes "Cut <a>.M (x).N \<longrightarrow>\<^sub>l R"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  2745
  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
  2746
  (\<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
  2747
  (\<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
  2748
                                                                            \<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
  2749
  (\<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
  2750
                                                                            \<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
  2751
  (\<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
  2752
                                                                            \<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
  2753
  (\<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
  2754
                                                                            \<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
  2755
  (\<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
  2756
            R = Cut <b>.(Cut <c>.N1 (z).M') (y).N2 \<and> b\<sharp>(c,N1) \<and> fic M a \<and> fin N x)"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2757
using assms
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2758
proof (cases rule: l_redu.cases)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2759
  case (LAxR x M a b)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2760
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2761
    apply(simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2762
    by (metis alpha(2) crename_rename)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2763
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2764
  case (LAxL a M x y)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2765
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2766
    apply(simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2767
    by (metis alpha(1) nrename_rename)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2768
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2769
  case (LNot y M N x a b)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2770
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2771
    apply(simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2772
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2773
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2774
    apply(rule disjI1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2775
    apply(erule conjE)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2776
    apply(generate_fresh "coname")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2777
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2778
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2779
    apply(drule_tac c="c" in  alpha_coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2780
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2781
    apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2782
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2783
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2784
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2785
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2786
    apply(simp add: calc_atm abs_fresh fresh_prod fresh_atm fresh_left)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2787
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2788
    apply(drule_tac z="ca" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2789
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2790
    apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2791
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2792
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2793
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2794
    apply(auto simp: calc_atm abs_fresh fresh_left)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2795
    using nrename_fresh nrename_swap apply presburger
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2796
    using crename_fresh crename_swap by presburger
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2797
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2798
  case (LAnd1 b a1 M1 a2 M2 N y x)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2799
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2800
    apply -
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2801
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2802
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2803
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2804
    apply(rule disjI1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2805
    apply(clarsimp simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2806
    apply(generate_fresh "coname")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2807
    apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2808
    apply(drule_tac c="c" in  alpha_coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2809
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2810
    apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2811
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2812
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2813
     apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2814
     apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2815
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2816
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2817
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2818
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2819
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2820
       apply(drule_tac z="ca" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2821
        apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2822
       apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2823
       apply (metis swap_simps(6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2824
      apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2825
      apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2826
      apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2827
      apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2828
       apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2829
      apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2830
      apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2831
      apply (smt (verit, del_insts) abs_fresh(1,2) abs_perm(1,2) fic.intros(3) fin.intros(3) fresh_bij(1) perm_fresh_fresh(1,2) swap_simps(1,6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2832
     apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2833
     apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2834
     apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2835
     apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2836
     apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2837
      apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2838
     apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2839
     apply (smt (verit) abs_fresh(1,2) abs_perm(1,2) fic.intros(3) fin.intros(3) perm_fresh_fresh(1,2) perm_swap(3) swap_simps(1,6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2840
    apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2841
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2842
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2843
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2844
    apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2845
     apply(simp add: fresh_prod fresh_atm abs_fresh cong: conj_cong)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2846
    apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2847
    by (smt (verit, best) abs_fresh(1,2) abs_perm(1) alpha(2) fic.intros(3) fin.intros(3) fresh_bij(1,2) perm_fresh_fresh(1,2) swap_simps(2,3))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2848
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2849
  case (LAnd2 b a1 M1 a2 M2 N y x)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2850
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2851
    apply -
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2852
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2853
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2854
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2855
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2856
    apply(rule disjI1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2857
    apply(clarsimp simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2858
    apply(generate_fresh "coname")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2859
    apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2860
    apply(drule_tac c="c" in  alpha_coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2861
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2862
    apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2863
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2864
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2865
     apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2866
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2867
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2868
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2869
    apply(auto simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2870
       apply(drule_tac z="ca" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2871
        apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2872
       apply (metis swap_simps(6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2873
      apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2874
      apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2875
      apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2876
      apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2877
       apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2878
      apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2879
      apply (smt (verit, ccfv_threshold) abs_fresh(1,2) abs_perm(1) fic.intros(3) fin.intros(4) fresh_bij(1,2) perm_fresh_fresh(1,2) swap_simps(1,4,6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2880
     apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2881
     apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2882
     apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2883
     apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2884
      apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2885
     apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2886
     apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2887
     apply (smt (verit) abs_fresh(1,2) abs_perm(1,2) fic.intros(3) fin.intros(4) fresh_bij(1) perm_fresh_fresh(1,2) swap_simps(1,6))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2888
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2889
    apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2890
    apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2891
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2892
    apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2893
    by (smt (verit, ccfv_SIG) abs_fresh(1,2) abs_perm(1) alpha(2) fic.intros(3) fin.intros(4) fresh_bij(1,2) perm_fresh_fresh(1,2) swap_simps(2,3))
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2894
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2895
  case (LOr1 b a M N1 N2 y x1 x2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2896
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2897
    apply -
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2898
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2899
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2900
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2901
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2902
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2903
    apply(rule disjI1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2904
    apply(clarsimp simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2905
    apply(generate_fresh "coname")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2906
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2907
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2908
    apply(drule_tac c="c" in  alpha_coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2909
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2910
    apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2911
    apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2912
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2913
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2914
     apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2915
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2916
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2917
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2918
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2919
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2920
     apply(drule_tac z="ca" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2921
      apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2922
     apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2923
     apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2924
     apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2925
      apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2926
      apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2927
       apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2928
      apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2929
     apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2930
        apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2931
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2932
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2933
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2934
    apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2935
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2936
    apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2937
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2938
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2939
     apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2940
     apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2941
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2942
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2943
    apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2944
     apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2945
    done
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2946
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2947
  case (LOr2 b a M N1 N2 y x1 x2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2948
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2949
    apply -
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2950
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2951
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2952
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2953
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2954
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2955
    apply(rule disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2956
    apply(rule disjI1)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2957
    apply(simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2958
    apply(erule conjE)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2959
    apply(generate_fresh "coname")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2960
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2961
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2962
    apply(drule_tac c="c" in  alpha_coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2963
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2964
    apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2965
     apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2966
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2967
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2968
     apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2969
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2970
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2971
     apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2972
     apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2973
     apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2974
      apply(drule_tac z="ca" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2975
       apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2976
      apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2977
      apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2978
      apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2979
      apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2980
       apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2981
      apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2982
      apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2983
     apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2984
     apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2985
     apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2986
     apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2987
     apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2988
      apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2989
     apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2990
     apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2991
     apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2992
     apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2993
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2994
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2995
     apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2996
       apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2997
    done
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2998
next
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  2999
  case (LImp z N y P x M b a c)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3000
  then show ?thesis
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3001
    apply(intro disjI2)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3002
    apply(clarsimp simp add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3003
    apply(generate_fresh "coname")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3004
    apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3005
    apply(drule_tac c="ca" in  alpha_coname)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3006
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3007
    apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3008
    apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3009
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3010
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3011
     apply(rule_tac s="a" and t="[(a,ca)]\<bullet>[(b,ca)]\<bullet>b" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3012
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3013
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3014
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3015
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3016
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3017
     apply(drule_tac z="caa" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3018
      apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3019
     apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3020
     apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3021
     apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3022
     apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3023
      apply(rule_tac s="x" and t="[(x,caa)]\<bullet>[(z,caa)]\<bullet>z" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3024
       apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3025
      apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3026
     apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3027
       apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3028
    apply(generate_fresh "name")
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3029
    apply(simp add: abs_fresh fresh_prod fresh_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3030
    apply(auto)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3031
    apply(drule_tac z="cb" in  alpha_name)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3032
     apply(simp add: fresh_prod fresh_atm abs_fresh)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3033
    apply(simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3034
    apply(perm_simp)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3035
    apply(rule exI)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3036
    apply(rule conjI)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3037
     apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(z,cb)]\<bullet>z" in subst)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3038
      apply(simp add: calc_atm)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3039
     apply(rule refl)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3040
    apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3041
     apply(perm_simp)+
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3042
    done
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3043
qed
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3044
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
inductive
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
  3047
  c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>c _\<close> [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3048
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3049
  left[intro]:  "\<lbrakk>\<not>fic M a; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c M{a:=(x).N}"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3050
| right[intro]: "\<lbrakk>\<not>fin N x; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c N{x:=<a>.M}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3051
9be4ab2acc13 split Class.thy into parts to 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
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
  3053
9be4ab2acc13 split Class.thy into parts to 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
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
  3055
 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
  3056
9be4ab2acc13 split Class.thy into parts to 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
lemma better_left[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3058
  shows "\<not>fic M a \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c M{a:=(x).N}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3059
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
  3060
  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
  3061
  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
  3062
  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
  3063
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3064
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3065
  also have "\<dots> \<longrightarrow>\<^sub>c ([(a',a)]\<bullet>M){a':=(x').([(x',x)]\<bullet>N)}" using fs1 fs2 not_fic
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3066
    apply(intro left)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3067
    apply (metis fic_rename perm_swap(4))
36277
9be4ab2acc13 split Class.thy into parts to 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(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
  3069
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  3071
    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
  3072
  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
  3073
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
  3074
9be4ab2acc13 split Class.thy into parts to 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
lemma better_right[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3076
  shows "\<not>fin N x \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c N{x:=<a>.M}"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3077
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
  3078
  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
  3079
  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
  3080
  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
  3081
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3082
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3083
  also have "\<dots> \<longrightarrow>\<^sub>c ([(x',x)]\<bullet>N){x':=<a'>.([(a',a)]\<bullet>M)}" using fs1 fs2 not_fin
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3084
    apply (intro right)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3085
    apply (metis fin_rename perm_swap(3))
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3086
    apply (simp add: fresh_left calc_atm)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3087
    done
9be4ab2acc13 split Class.thy into parts to 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
  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
  3089
    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
  3090
  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
  3091
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
  3092
9be4ab2acc13 split Class.thy into parts to 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
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
  3094
  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
  3095
  and   c::"coname"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3096
  shows "M \<longrightarrow>\<^sub>c M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3097
  and   "M \<longrightarrow>\<^sub>c M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3098
  by (induct rule: c_redu.induct) (auto simp: abs_fresh rename_fresh subst_fresh)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3099
9be4ab2acc13 split Class.thy into parts to 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
inductive
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
  3101
  a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>a _\<close> [100,100] 100)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3102
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3103
  al_redu[intro]: "M\<longrightarrow>\<^sub>l M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3104
| ac_redu[intro]: "M\<longrightarrow>\<^sub>c M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3105
| a_Cut_l: "\<lbrakk>a\<sharp>N; x\<sharp>M; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M' (x).N"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3106
| a_Cut_r: "\<lbrakk>a\<sharp>N; x\<sharp>M; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M (x).N'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3107
| a_NotL[intro]: "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^sub>a NotL <a>.M' x"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3108
| a_NotR[intro]: "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^sub>a NotR (x).M' a"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3109
| a_AndR_l: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M' <b>.N c"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3110
| a_AndR_r: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M <b>.N' c"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3111
| a_AndL1: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a AndL1 (x).M' y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3112
| a_AndL2: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a AndL2 (x).M' y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3113
| a_OrL_l: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M' (y).N z"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3114
| a_OrL_r: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M (y).N' z"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3115
| a_OrR1: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a OrR1 <a>.M' b"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3116
| a_OrR2: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a OrR2 <a>.M' b"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3117
| a_ImpL_l: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M' (x).N y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3118
| a_ImpL_r: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M (x).N' y"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3119
| a_ImpR: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a ImpR (x).<a>.M' b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3120
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3121
lemma fresh_a_redu1:
36277
9be4ab2acc13 split Class.thy into parts to 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
  fixes x::"name"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3123
  shows "M \<longrightarrow>\<^sub>a M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3124
  by (induct rule: a_redu.induct) (auto simp: fresh_l_redu fresh_c_redu abs_fresh abs_supp fin_supp)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3125
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3126
lemma fresh_a_redu2:
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3127
  fixes c::"coname"
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3128
  shows "M \<longrightarrow>\<^sub>a M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3129
  by (induct rule: a_redu.induct) (auto simp: fresh_l_redu fresh_c_redu abs_fresh abs_supp fin_supp)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3130
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3131
lemmas fresh_a_redu = fresh_a_redu1 fresh_a_redu2
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  3134
9be4ab2acc13 split Class.thy into parts to 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
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
  3136
  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
  3137
9be4ab2acc13 split Class.thy into parts to 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
lemma better_CutL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3139
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M' (x).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3140
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3141
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3142
  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
  3143
  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
  3144
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3145
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3146
  also have "\<dots> \<longrightarrow>\<^sub>a  Cut <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N)" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3147
    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
  3148
  also have "\<dots> = Cut <a>.M' (x).N" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3149
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3150
  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
  3151
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
  3152
9be4ab2acc13 split Class.thy into parts to 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
lemma better_CutR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3154
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M (x).N'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3155
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3156
  assume red: "N\<longrightarrow>\<^sub>a N'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3157
  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
  3158
  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
  3159
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3160
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3161
  also have "\<dots> \<longrightarrow>\<^sub>a  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N')" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3162
    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
  3163
  also have "\<dots> = Cut <a>.M (x).N'" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3164
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3165
  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
  3166
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
  3167
    
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndRL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3169
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M' <b>.N c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3170
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3171
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3172
  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
  3173
  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
  3174
  have "AndR <a>.M <b>.N c =  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3175
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3176
  also have "\<dots> \<longrightarrow>\<^sub>a  AndR <a'>.([(a',a)]\<bullet>M') <b'>.([(b',b)]\<bullet>N) c" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3177
    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
  3178
  also have "\<dots> = AndR <a>.M' <b>.N c" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3179
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3180
  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
  3181
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
  3182
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndRR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3184
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M <b>.N' c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3185
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3186
  assume red: "N\<longrightarrow>\<^sub>a N'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3187
  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
  3188
  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
  3189
  have "AndR <a>.M <b>.N c =  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3190
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3191
  also have "\<dots> \<longrightarrow>\<^sub>a  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N') c" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3192
    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
  3193
  also have "\<dots> = AndR <a>.M <b>.N' c" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3194
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3195
  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
  3196
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
  3197
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndL1_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3199
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a AndL1 (x).M' y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3200
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3201
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3202
  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
  3203
  have "AndL1 (x).M y = AndL1 (x').([(x',x)]\<bullet>M) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3204
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3205
  also have "\<dots> \<longrightarrow>\<^sub>a AndL1 (x').([(x',x)]\<bullet>M') y" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3206
    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
  3207
  also have "\<dots> = AndL1 (x).M' y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3208
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3209
  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
  3210
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
  3211
9be4ab2acc13 split Class.thy into parts to 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
lemma better_AndL2_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3213
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a AndL2 (x).M' y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3214
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3215
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3216
  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
  3217
  have "AndL2 (x).M y = AndL2 (x').([(x',x)]\<bullet>M) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3218
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3219
  also have "\<dots> \<longrightarrow>\<^sub>a AndL2 (x').([(x',x)]\<bullet>M') y" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3220
    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
  3221
  also have "\<dots> = AndL2 (x).M' y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3222
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3223
  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
  3224
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
  3225
9be4ab2acc13 split Class.thy into parts to 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
lemma better_OrLL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3227
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M' (y).N z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3228
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3229
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3230
  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
  3231
  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
  3232
  have "OrL (x).M (y).N z =  OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3233
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3234
  also have "\<dots> \<longrightarrow>\<^sub>a OrL (x').([(x',x)]\<bullet>M') (y').([(y',y)]\<bullet>N) z" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3235
    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
  3236
  also have "\<dots> = OrL (x).M' (y).N z" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3237
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3238
  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
  3239
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
  3240
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3241
lemma better_OrLR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3242
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M (y).N' z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3243
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3244
  assume red: "N\<longrightarrow>\<^sub>a N'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3245
  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
  3246
  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
  3247
  have "OrL (x).M (y).N z =  OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3248
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3249
  also have "\<dots> \<longrightarrow>\<^sub>a OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N') z" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3250
    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
  3251
  also have "\<dots> = OrL (x).M (y).N' z" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3252
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3253
  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
  3254
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
  3255
9be4ab2acc13 split Class.thy into parts to 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
lemma better_OrR1_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3257
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a OrR1 <a>.M' b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3258
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3259
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3260
  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
  3261
  have "OrR1 <a>.M b = OrR1 <a'>.([(a',a)]\<bullet>M) b"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3262
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3263
  also have "\<dots> \<longrightarrow>\<^sub>a OrR1 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3264
    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
  3265
  also have "\<dots> = OrR1 <a>.M' b" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3266
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3267
  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
  3268
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
  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 better_OrR2_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3271
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a OrR2 <a>.M' b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3272
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3273
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3274
  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
  3275
  have "OrR2 <a>.M b = OrR2 <a'>.([(a',a)]\<bullet>M) b"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3276
    using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3277
  also have "\<dots> \<longrightarrow>\<^sub>a OrR2 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3278
    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
  3279
  also have "\<dots> = OrR2 <a>.M' b" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3280
    using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3281
  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
  3282
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
  3283
9be4ab2acc13 split Class.thy into parts to 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
lemma better_ImpLL_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3285
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M' (x).N y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3286
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3287
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3288
  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
  3289
  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
  3290
  have "ImpL <a>.M (x).N y =  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3291
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3292
  also have "\<dots> \<longrightarrow>\<^sub>a  ImpL <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N) y" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3293
    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
  3294
  also have "\<dots> = ImpL <a>.M' (x).N y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3295
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3296
  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
  3297
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
  3298
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3299
lemma better_ImpLR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3300
  shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M (x).N' y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3301
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3302
  assume red: "N\<longrightarrow>\<^sub>a N'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3303
  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
  3304
  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
  3305
  have "ImpL <a>.M (x).N y =  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3306
    using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3307
  also have "\<dots> \<longrightarrow>\<^sub>a  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N') y" using fs1 fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3308
    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
  3309
  also have "\<dots> = ImpL <a>.M (x).N' y" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3310
    using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3311
  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
  3312
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
  3313
9be4ab2acc13 split Class.thy into parts to 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
lemma better_ImpR_intro[intro]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3315
  shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a ImpR (x).<a>.M' b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3316
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3317
  assume red: "M\<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3318
  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
  3319
  have "ImpR (x).<a>.M b = ImpR (x).<a'>.([(a',a)]\<bullet>M) b"
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3320
    using fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3321
  also have "\<dots> \<longrightarrow>\<^sub>a ImpR (x).<a'>.([(a',a)]\<bullet>M') b" using fs2 red
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3322
    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
  3323
  also have "\<dots> = ImpR (x).<a>.M' b" 
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3324
    using fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3325
  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
  3326
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
  3327
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  3328
text \<open>axioms do not reduce\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3329
9be4ab2acc13 split Class.thy into parts to 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
lemma ax_do_not_l_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3331
  shows "Ax x a \<longrightarrow>\<^sub>l M \<Longrightarrow> False"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3332
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
  3333
 
9be4ab2acc13 split Class.thy into parts to 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
lemma ax_do_not_c_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3335
  shows "Ax x a \<longrightarrow>\<^sub>c M \<Longrightarrow> False"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3336
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
  3337
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3338
lemma ax_do_not_a_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3339
  shows "Ax x a \<longrightarrow>\<^sub>a M \<Longrightarrow> False"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3340
apply(erule_tac a_redu.cases) 
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3341
apply(simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3342
  using ax_do_not_l_reduce ax_do_not_c_reduce by blast+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3343
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_NotL_elim:
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3345
  assumes "NotL <a>.M x \<longrightarrow>\<^sub>a R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3346
  shows "\<exists>M'. R = NotL <a>.M' x \<and> M\<longrightarrow>\<^sub>aM'"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3347
  using assms
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3348
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3349
    apply(erule_tac l_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3350
   apply(erule_tac c_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3351
  apply (smt (verit, best) a_redu.eqvt(2) alpha(2) fresh_a_redu2)+
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3352
  done
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3353
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_NotR_elim:
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3355
  assumes "NotR (x).M a \<longrightarrow>\<^sub>a R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3356
  shows "\<exists>M'. R = NotR (x).M' a \<and> M\<longrightarrow>\<^sub>aM'"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3357
  using assms
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3358
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3359
    apply(erule_tac l_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3360
   apply(erule_tac c_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3361
  apply (smt (verit, best) a_redu.eqvt(1) alpha(1) fresh_a_redu1)+
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3362
  done
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3363
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_AndR_elim:
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3365
  assumes "AndR <a>.M <b>.N c\<longrightarrow>\<^sub>a R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3366
  shows "(\<exists>M'. R = AndR <a>.M' <b>.N c \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = AndR <a>.M <b>.N' c \<and> N\<longrightarrow>\<^sub>aN')"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3367
  using assms
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3368
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3369
     apply(erule_tac l_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3370
    apply(erule_tac c_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3371
   apply (smt (verit) a_redu.eqvt(2) alpha'(2) fresh_a_redu2)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3372
  by (metis a_NotL a_redu_NotL_elim trm.inject(4))
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_AndL1_elim:
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3375
  assumes "AndL1 (x).M y \<longrightarrow>\<^sub>a R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3376
  shows "\<exists>M'. R = AndL1 (x).M' y \<and> M\<longrightarrow>\<^sub>aM'"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3377
  using assms
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3378
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3379
    apply(erule_tac l_redu.cases, simp_all add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3380
   apply(erule_tac c_redu.cases, simp_all add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3381
  by (metis a_NotR a_redu_NotR_elim trm.inject(3))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  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
lemma a_redu_AndL2_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3384
  assumes a: "AndL2 (x).M y \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3385
  shows "\<exists>M'. R = AndL2 (x).M' y \<and> M\<longrightarrow>\<^sub>aM'"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3386
  using a
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3387
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3388
  apply(erule_tac l_redu.cases, simp_all add: trm.inject)
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3389
   apply(erule_tac c_redu.cases, simp_all add: trm.inject)
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3390
  by (metis a_NotR a_redu_NotR_elim trm.inject(3))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3391
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_OrL_elim:
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3393
  assumes "OrL (x).M (y).N z\<longrightarrow>\<^sub>a R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3394
  shows "(\<exists>M'. R = OrL (x).M' (y).N z \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = OrL (x).M (y).N' z \<and> N\<longrightarrow>\<^sub>aN')"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3395
  using assms
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3396
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3397
     apply(erule_tac l_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3398
    apply(erule_tac c_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3399
   apply (metis a_NotR a_redu_NotR_elim trm.inject(3))+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3400
done
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_OrR1_elim:
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3403
  assumes "OrR1 <a>.M b \<longrightarrow>\<^sub>a R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3404
  shows "\<exists>M'. R = OrR1 <a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3405
  using assms
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3406
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3407
    apply(erule_tac l_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3408
   apply(erule_tac c_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3409
  by (metis a_NotL a_redu_NotL_elim trm.inject(4))
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_OrR2_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3412
  assumes a: "OrR2 <a>.M b \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3413
  shows "\<exists>M'. R = OrR2 <a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3414
  using assms
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3415
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3416
    apply(erule_tac l_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3417
   apply(erule_tac c_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3418
  by (metis a_redu_OrR1_elim better_OrR1_intro trm.inject(8))
36277
9be4ab2acc13 split Class.thy into parts to 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 a_redu_ImpL_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3421
  assumes a: "ImpL <a>.M (y).N z\<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3422
  shows "(\<exists>M'. R = ImpL <a>.M' (y).N z \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = ImpL <a>.M (y).N' z \<and> N\<longrightarrow>\<^sub>aN')"
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3423
  using assms
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3424
  apply(erule_tac a_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3425
     apply(erule_tac l_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3426
    apply(erule_tac c_redu.cases, simp_all add: trm.inject)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3427
   apply (smt (verit) a_redu.eqvt(2) alpha'(2) fresh_a_redu2)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3428
  by (metis a_NotR a_redu_NotR_elim trm.inject(3))
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3429
9be4ab2acc13 split Class.thy into parts to 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
lemma a_redu_ImpR_elim:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3431
  assumes a: "ImpR (x).<a>.M b \<longrightarrow>\<^sub>a R"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3432
  shows "\<exists>M'. R = ImpR (x).<a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
71989
bad75618fb82 extraction of equations x = t from premises beneath meta-all
haftmann
parents: 67613
diff changeset
  3433
using a [[simproc del: defined_all]]
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3434
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
  3435
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
  3436
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
  3437
apply(auto)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56073
diff changeset
  3438
apply(rotate_tac 3)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3439
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
  3440
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3441
  apply(erule_tac c_redu.cases, simp_all add: trm.inject)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3442
apply(auto simp: alpha a_redu.eqvt abs_perm abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3443
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3444
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3445
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3446
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3447
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3448
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3449
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3450
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3451
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3452
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3453
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3454
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3455
apply(rule_tac x="([(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3456
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3457
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
  3458
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3459
apply(rule 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
  3460
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
  3461
apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3462
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3463
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
  3464
apply(rule trans)
9be4ab2acc13 split Class.thy into parts to 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(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
  3466
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
  3467
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3468
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3469
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3470
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3471
apply(rule_tac x="([(a,aa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI)
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3472
apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
80611
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3473
  apply (simp add: cp_coname_name1 perm_dj(4) perm_swap(3,4))
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3474
  apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI)
fbc859ccdaf3 A massive reduction of some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents: 80609
diff changeset
  3475
by(simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap perm_compose(4) perm_dj(4))
36277
9be4ab2acc13 split Class.thy into parts to 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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  3477
text \<open>Transitive Closure\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3478
9be4ab2acc13 split Class.thy into parts to 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
abbreviation
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80651
diff changeset
  3480
 a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>a* _\<close> [80,80] 80)
36277
9be4ab2acc13 split Class.thy into parts to 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
where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 66453
diff changeset
  3482
  "M \<longrightarrow>\<^sub>a* M' \<equiv> (a_redu)\<^sup>*\<^sup>* M M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3483
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3484
lemmas a_starI = r_into_rtranclp
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_starE:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3487
  assumes a: "M \<longrightarrow>\<^sub>a* M'"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3488
  shows "M = M' \<or> (\<exists>N. M \<longrightarrow>\<^sub>a N \<and> N \<longrightarrow>\<^sub>a* M')"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3489
  using a  by (induct) (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
  3490
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_refl:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3492
  shows "M \<longrightarrow>\<^sub>a* M"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3493
  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
  3494
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3495
declare rtranclp_trans [trans]
36277
9be4ab2acc13 split Class.thy into parts to 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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  3497
text \<open>congruence rules for \<open>\<longrightarrow>\<^sub>a*\<close>\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3498
9be4ab2acc13 split Class.thy into parts to 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
lemma ax_do_not_a_star_reduce:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3500
  shows "Ax x a \<longrightarrow>\<^sub>a* M \<Longrightarrow> M = Ax x a"
80609
4b5d3d0abb69 More simplification of proofs. Trying to fix the syntax too
paulson <lp15@cam.ac.uk>
parents: 80595
diff changeset
  3501
  using a_starE ax_do_not_a_reduce by blast
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3502
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_CutL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3505
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M' (x).N"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3506
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
  3507
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_CutR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3509
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M (x).N'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3510
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
  3511
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_Cut:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3513
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M' (x).N'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3514
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
  3515
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_NotR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3517
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^sub>a* NotR (x).M' a"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3518
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
  3519
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_NotL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3521
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^sub>a* NotL <a>.M' x"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3522
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
  3523
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndRL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3525
    "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M' <b>.N c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3526
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
  3527
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndRR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3529
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M <b>.N' c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3530
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
  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 a_star_AndR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3533
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M' <b>.N' c"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3534
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
  3535
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndL1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3537
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a* AndL1 (x).M' y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3538
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
  3539
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_AndL2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3541
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a* AndL2 (x).M' y"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3542
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
  3543
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrLL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3545
    "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M' (y).N z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3546
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
  3547
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrLR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3549
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M (y).N' z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3550
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
  3551
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3553
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M' (y).N' z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3554
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
  3555
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrR1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3557
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a* OrR1 <a>.M' b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3558
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
  3559
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_OrR2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3561
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a* OrR2 <a>.M' b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3562
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
  3563
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_ImpLL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3565
    "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M' (y).N z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3566
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
  3567
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_ImpLR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3569
    "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M (y).N' z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3570
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
  3571
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3572
lemma a_star_ImpL:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3573
    "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M' (y).N' z"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3574
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
  3575
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_ImpR:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3577
    "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a* ImpR (x).<a>.M' b"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3578
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
  3579
9be4ab2acc13 split Class.thy into parts to 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
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
  3581
                      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
  3582
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_NotL_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3584
  assumes "NotL <a>.M x \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3585
  shows "\<exists>M'. R = NotL <a>.M' x \<and> M \<longrightarrow>\<^sub>a* M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3586
  using assms by (induct set: rtranclp) (use a_redu_NotL_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3587
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_NotR_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3589
  assumes "NotR (x).M a \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3590
  shows "\<exists>M'. R = NotR (x).M' a \<and> M \<longrightarrow>\<^sub>a* M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3591
  using assms by (induct set: rtranclp) (use a_redu_NotR_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3592
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_AndR_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3594
  assumes "AndR <a>.M <b>.N c\<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3595
  shows "(\<exists>M' N'. R = AndR <a>.M' <b>.N' c \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<longrightarrow>\<^sub>a* N')"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3596
  using assms by (induct set: rtranclp) (use a_redu_AndR_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3597
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_AndL1_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3599
  assumes "AndL1 (x).M y \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3600
  shows "\<exists>M'. R = AndL1 (x).M' y \<and> M \<longrightarrow>\<^sub>a* M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3601
  using assms by (induct set: rtranclp) (use a_redu_AndL1_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3602
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_AndL2_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3604
  assumes "AndL2 (x).M y \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3605
  shows "\<exists>M'. R = AndL2 (x).M' y \<and> M \<longrightarrow>\<^sub>a* M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3606
  using assms by (induct set: rtranclp) (use a_redu_AndL2_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3607
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_OrL_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3609
  assumes "OrL (x).M (y).N z \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3610
  shows "(\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<longrightarrow>\<^sub>a* N')"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3611
  using assms by (induct set: rtranclp) (use a_redu_OrL_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3612
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_OrR1_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3614
  assumes "OrR1 <a>.M y \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3615
  shows "\<exists>M'. R = OrR1 <a>.M' y \<and> M \<longrightarrow>\<^sub>a* M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3616
  using assms by (induct set: rtranclp) (use a_redu_OrR1_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3617
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_OrR2_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3619
  assumes "OrR2 <a>.M y \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3620
  shows "\<exists>M'. R = OrR2 <a>.M' y \<and> M \<longrightarrow>\<^sub>a* M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3621
  using assms by (induct set: rtranclp) (use a_redu_OrR2_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3622
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_ImpR_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3624
  assumes "ImpR (x).<a>.M y \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3625
  shows "\<exists>M'. R = ImpR (x).<a>.M' y \<and> M \<longrightarrow>\<^sub>a* M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3626
  using assms by (induct set: rtranclp) (use a_redu_ImpR_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3627
9be4ab2acc13 split Class.thy into parts to 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
lemma a_star_redu_ImpL_elim:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3629
  assumes "ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3630
  shows "(\<exists>M' N'. R = ImpL <a>.M' (y).N' z \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<longrightarrow>\<^sub>a* N')"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3631
  using assms by (induct set: rtranclp) (use a_redu_ImpL_elim in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3632
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  3633
text \<open>Substitution\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3634
9be4ab2acc13 split Class.thy into parts to 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
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
  3636
  shows "\<not>fin(M{x:=<c>.P}) x"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3637
proof (nominal_induct M avoiding: x c P rule: trm.strong_induct)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3638
  case (Ax name coname)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3639
  with fin_rest_elims show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3640
    by (auto simp: fin_Ax_elim)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3641
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3642
  case (NotL coname trm name)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3643
  obtain x'::name where "x'\<sharp>(trm{x:=<c>.P},P)"
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3644
    by (meson exists_fresh(1) fs_name1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3645
  with NotL fin_NotL_elim fresh_fun_simp_NotL show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3646
    by simp (metis fin_rest_elims(1) fresh_prod)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3647
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3648
  case (AndL1 name1 trm name2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3649
  obtain x'::name where "x' \<sharp> (trm{x:=<c>.P},P,name1)"
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3650
    by (meson exists_fresh(1) fs_name1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3651
  with AndL1 fin_AndL1_elim fresh_fun_simp_AndL1 show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3652
    by simp (metis fin_rest_elims(1) fresh_prod)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3653
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3654
  case (AndL2 name2 trm name2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3655
  obtain x'::name where "x' \<sharp> (trm{x:=<c>.P},P,name2)"
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3656
    by (meson exists_fresh(1) fs_name1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3657
  with AndL2 fin_AndL2_elim fresh_fun_simp_AndL2 better_AndL2_substn show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3658
    by (metis abs_fresh(1) fresh_atm(1) not_fin_subst2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3659
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3660
  case (OrL name1 trm1 name2 trm2 name3)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3661
  obtain x'::name where "x' \<sharp> (trm1{x:=<c>.P},P,name1,trm2{x:=<c>.P},name2)"
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3662
    by (meson exists_fresh(1) fs_name1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3663
  with OrL fin_OrL_elim fresh_fun_simp_OrL show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3664
    by simp (metis fin_rest_elims(1) fresh_prod)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3665
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3666
  case (ImpL coname trm1 name1 trm2 name2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3667
  obtain x'::name where "x' \<sharp> (trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})"
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3668
    by (meson exists_fresh(1) fs_name1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3669
  with ImpL fin_ImpL_elim fresh_fun_simp_ImpL show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3670
    by simp (metis fin_rest_elims(1) fresh_prod)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3671
qed (use fin_rest_elims not_fin_subst2 in auto)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3672
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3673
lemmas subst_not_fin2 = not_fin_subst1
36277
9be4ab2acc13 split Class.thy into parts to 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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  3675
text \<open>Reductions\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3676
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_l_reduce:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3678
  assumes "fin M x" and "M \<longrightarrow>\<^sub>l M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3679
  shows "fin M' x"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3680
  using assms fin_rest_elims(1) l_redu.simps by fastforce
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
lemma fin_c_reduce:
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3683
  assumes "fin M x" and "M \<longrightarrow>\<^sub>c M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3684
  shows "fin M' x"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3685
  using assms c_redu.simps fin_rest_elims(1) by fastforce
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3686
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3687
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
  3688
  assumes  a: "fin M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3689
  and      b: "M \<longrightarrow>\<^sub>a M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3690
  shows "fin M' x"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3691
using assms
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3692
proof (induct)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3693
  case (1 x a)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3694
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3695
    using ax_do_not_a_reduce by auto
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3696
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3697
  case (2 x M a)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3698
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3699
    using a_redu_NotL_elim fresh_a_redu1 by blast
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3700
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3701
  case (3 y x M)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3702
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3703
    by (metis a_redu_AndL1_elim abs_fresh(1) fin.intros(3) fresh_a_redu1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3704
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3705
  case (4 y x M)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3706
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3707
    by (metis a_redu_AndL2_elim abs_fresh(1) fin.intros(4) fresh_a_redu1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3708
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3709
  case (5 z x M y N)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3710
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3711
    by (smt (verit) a_redu_OrL_elim abs_fresh(1) fin.intros(5) fresh_a_redu1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3712
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3713
  case (6 y M x N a)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3714
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3715
    by (smt (verit, best) a_redu_ImpL_elim abs_fresh(1) fin.simps fresh_a_redu1)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3716
qed
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3717
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  3720
  assumes  a: "fin M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3721
  and      b: "M \<longrightarrow>\<^sub>a* M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3722
  shows "fin M' x"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3723
using b a by (induct set: rtranclp)(auto simp: fin_a_reduce)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3724
9be4ab2acc13 split Class.thy into parts to 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
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
  3726
  assumes  a: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3727
  and      b: "M \<longrightarrow>\<^sub>l M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3728
  shows "fic M' x"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3729
  using b a
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3730
  by induction (use fic_rest_elims in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3731
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3732
lemma 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
  3733
  assumes a: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3734
  and     b: "M \<longrightarrow>\<^sub>c M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3735
  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
  3736
using b a
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3737
  by induction (use fic_rest_elims in force)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3738
9be4ab2acc13 split Class.thy into parts to 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
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
  3740
  assumes a: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3741
  and     b: "M \<longrightarrow>\<^sub>a M'"
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3742
shows "fic M' x"
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3743
  using assms
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3744
proof induction
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3745
  case (1 x a)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3746
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3747
    using ax_do_not_a_reduce by fastforce
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3748
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3749
  case (2 a M x)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3750
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3751
    using a_redu_NotR_elim fresh_a_redu2 by blast
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3752
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3753
  case (3 c a M b N)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3754
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3755
    by (smt (verit) a_redu_AndR_elim abs_fresh(2) fic.intros(3) fresh_a_redu2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3756
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3757
  case (4 b a M)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3758
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3759
    by (metis a_redu_OrR1_elim abs_fresh(2) fic.intros(4) fresh_a_redu2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3760
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3761
  case (5 b a M)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3762
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3763
    by (metis a_redu_OrR2_elim abs_fresh(2) fic.simps fresh_a_redu2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3764
next
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3765
  case (6 b a M x)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3766
  then show ?case
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3767
    by (metis a_redu_ImpR_elim abs_fresh(2) fic.simps fresh_a_redu2)
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3768
qed
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3769
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  3772
  assumes  a: "fic M x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3773
  and      b: "M \<longrightarrow>\<^sub>a* M'"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3774
  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
  3775
using b a
80614
21bb6d17d58e Adjusting the precedences to reduce syntactic ambiguity
paulson <lp15@cam.ac.uk>
parents: 80611
diff changeset
  3776
  by (induct set: rtranclp) (auto simp: fic_a_reduce)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3777
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  3778
text \<open>substitution properties\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3779
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_with_ax1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3781
  shows "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3782
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
  3783
  case (Ax z b x a y)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3784
  show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3785
  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
  3786
    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
  3787
    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
  3788
    have "(Ax z b){x:=<a>.Ax y a} = Cut <a>.Ax y a (x).Ax x b" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3789
    also have "\<dots> \<longrightarrow>\<^sub>a* (Ax x b)[x\<turnstile>n>y]" by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3790
    finally show "Ax z b{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]" using eq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3791
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  3793
    assume neq: "z\<noteq>x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3794
    then show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]" using neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3795
  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
  3796
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3798
  have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>N" "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>M" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3799
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3800
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3801
  show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3802
  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
  3803
    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
  3804
    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
  3805
    have "(Cut <b>.M (z).N){x:=<a>.Ax y a} = Cut <a>.Ax y a (z).(N{x:=<a>.Ax y a})" using fs eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3806
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.Ax y a (z).(N[x\<turnstile>n>y])" using ih2 a_star_congs by blast
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3807
    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
  3808
      by (simp add: trm.inject alpha calc_atm fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3809
    finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3810
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  3812
    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
  3813
    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
  3814
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3815
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using ih1 ih2 a_star_congs by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3816
    finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3817
  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
  3818
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3820
  have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>b" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3821
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3822
  have "(NotR (z).M b){x:=<a>.Ax y a} = NotR (z).(M{x:=<a>.Ax y a}) b" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3823
  also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[x\<turnstile>n>y]) b" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3824
  finally show "(NotR (z).M b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotR (z).M b)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3825
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3827
  have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3828
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3829
  show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3830
  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
  3831
    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
  3832
    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
  3833
    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
  3834
    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
  3835
                        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
  3836
      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
  3837
    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
  3838
      using new by (simp add: fresh_fun_simp_NotL fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3839
    also have "\<dots> \<longrightarrow>\<^sub>a* (NotL <b>.(M{x:=<a>.Ax y a}) x')[x'\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3840
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  3841
      by (intro a_starI al_redu better_LAxL_intro) 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
  3842
    also have "\<dots> = NotL <b>.(M{x:=<a>.Ax y a}) y" using new by (simp add: nrename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3843
    also have "\<dots> \<longrightarrow>\<^sub>a* NotL <b>.(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3844
    also have "\<dots> = (NotL <b>.M z)[x\<turnstile>n>y]" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3845
    finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3846
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  3848
    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
  3849
    have "(NotL <b>.M z){x:=<a>.Ax y a} = NotL <b>.(M{x:=<a>.Ax y a}) z" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3850
    also have "\<dots> \<longrightarrow>\<^sub>a* NotL <b>.(M[x\<turnstile>n>y]) z" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3851
    finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]" using neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3852
  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
  3853
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3855
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "d\<sharp>x" "d\<sharp>a" "d\<sharp>y" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3856
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3857
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3858
  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
  3859
    using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3860
  also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[x\<turnstile>n>y]) <d>.(N[x\<turnstile>n>y]) e" using ih1 ih2 by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3861
  finally show "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3862
    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
  3863
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3865
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3866
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3867
  show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3868
  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
  3869
    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
  3870
    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
  3871
    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
  3872
    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
  3873
                        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
  3874
      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
  3875
    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
  3876
      using new by (simp add: fresh_fun_simp_AndL1 fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3877
    also have "\<dots> \<longrightarrow>\<^sub>a* (AndL1 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3878
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  3879
      by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3880
    also have "\<dots> = AndL1 (u).(M{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3881
      by (auto simp: fresh_prod fresh_atm nrename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3882
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3883
    also have "\<dots> = (AndL1 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3884
    finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3885
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  3887
    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
  3888
    have "(AndL1 (u).M v){x:=<a>.Ax y a} = AndL1 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3889
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3890
    finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3891
  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
  3892
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3894
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3895
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3896
  show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3897
  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
  3898
    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
  3899
    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
  3900
    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
  3901
    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
  3902
                        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
  3903
      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
  3904
    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
  3905
      using new by (simp add: fresh_fun_simp_AndL2 fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3906
    also have "\<dots> \<longrightarrow>\<^sub>a* (AndL2 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3907
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  3908
      by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3909
    also have "\<dots> = AndL2 (u).(M{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3910
      by (auto simp: fresh_prod fresh_atm nrename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3911
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3912
    also have "\<dots> = (AndL2 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3913
    finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3914
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  3916
    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
  3917
    have "(AndL2 (u).M v){x:=<a>.Ax y a} = AndL2 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3918
    also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3919
    finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3920
  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
  3921
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3923
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3924
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3925
  have "(OrR1 <c>.M d){x:=<a>.Ax y a} = OrR1 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3926
  also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3927
  finally show "(OrR1 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3928
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3930
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3931
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3932
  have "(OrR2 <c>.M d){x:=<a>.Ax y a} = OrR2 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3933
  also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3934
  finally show "(OrR2 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3935
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3937
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "v\<sharp>x" "v\<sharp>a" "v\<sharp>y" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3938
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3939
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3940
  show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3941
  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
  3942
    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
  3943
    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
  3944
    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
  3945
      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
  3946
    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
  3947
                 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
  3948
      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
  3949
    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
  3950
      using new by (simp add: fresh_fun_simp_OrL)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3951
    also have "\<dots> \<longrightarrow>\<^sub>a* (OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z')[z'\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3952
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  3953
      by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp_all add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3954
    also have "\<dots> = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3955
      by (auto simp: fresh_prod fresh_atm nrename_fresh subst_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3956
    also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) y" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3957
      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
  3958
    also have "\<dots> = (OrL (u).M (v).N z)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3959
    finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3960
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  3962
    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
  3963
    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
  3964
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3965
    also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) z" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3966
      using ih1 ih2 by (auto intro: a_star_congs)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3967
    finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3968
  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
  3969
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3971
  have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "z\<sharp>d" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3972
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3973
  have "(ImpR (z).<c>.M d){x:=<a>.Ax y a} = ImpR (z).<c>.(M{x:=<a>.Ax y a}) d" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3974
  also have "\<dots> \<longrightarrow>\<^sub>a* ImpR (z).<c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3975
  finally show "(ImpR (z).<c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3976
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  3978
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3979
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3980
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3981
  show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3982
  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
  3983
    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
  3984
    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
  3985
    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
  3986
      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
  3987
    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
  3988
                 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
  3989
      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
  3990
    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
  3991
      using new by (simp add: fresh_fun_simp_ImpL)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3992
    also have "\<dots> \<longrightarrow>\<^sub>a* (ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3993
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  3994
      by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp_all add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3995
    also have "\<dots> = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) y" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  3996
      by (auto simp: fresh_prod subst_fresh fresh_atm trm.inject alpha rename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  3997
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) y" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  3998
      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
  3999
    also have "\<dots> = (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4000
    finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4001
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4003
    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
  4004
    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
  4005
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4006
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) v" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4007
      using ih1 ih2 by (auto intro: a_star_congs)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4008
    finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4009
      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
  4010
  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
  4011
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
  4012
9be4ab2acc13 split Class.thy into parts to 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
lemma subst_with_ax2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4014
  shows "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4015
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
  4016
  case (Ax z c b a x)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4017
  show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4018
  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
  4019
    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
  4020
    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
  4021
    have "(Ax z c){b:=(x).Ax x a} = Cut <b>.Ax z c (x).Ax x a" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4022
    also have "\<dots> \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" using eq by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4023
    finally show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4024
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4026
    assume neq: "c\<noteq>b"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4027
    then show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4028
  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
  4029
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4031
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>N" "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>M" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4032
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4033
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4034
  show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4035
  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
  4036
    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
  4037
    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
  4038
    have "(Cut <c>.M (z).N){b:=(x).Ax x a} = Cut <c>.(M{b:=(x).Ax x a}) (x).Ax x a" using eq fs by simp 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4039
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <c>.(M[b\<turnstile>c>a]) (x).Ax x a" using ih1 a_star_congs by blast
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4040
    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
  4041
      by (simp add: trm.inject alpha calc_atm fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4042
    finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4043
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4045
    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
  4046
    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
  4047
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4048
    also have "\<dots> \<longrightarrow>\<^sub>a* Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using ih1 ih2 a_star_congs by blast
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4049
    finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4050
  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
  4051
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4053
  have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>c" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4054
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4055
  show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4056
  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
  4057
    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
  4058
    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
  4059
    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
  4060
    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
  4061
                        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
  4062
      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
  4063
    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
  4064
      using new by (simp add: fresh_fun_simp_NotR fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4065
    also have "\<dots> \<longrightarrow>\<^sub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4066
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4067
      by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) 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
  4068
    also have "\<dots> = NotR (z).(M{b:=(x).Ax x a}) a" using new by (simp add: crename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4069
    also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4070
    also have "\<dots> = (NotR (z).M c)[b\<turnstile>c>a]" using eq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4071
    finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4072
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4074
    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
  4075
    have "(NotR (z).M c){b:=(x).Ax x a} = NotR (z).(M{b:=(x).Ax x a}) c" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4076
    also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[b\<turnstile>c>a]) c" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4077
    finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]" using neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4078
  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
  4079
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4081
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4082
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4083
  have "(NotL <c>.M z){b:=(x).Ax x a} = NotL <c>.(M{b:=(x).Ax x a}) z" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4084
  also have "\<dots> \<longrightarrow>\<^sub>a* NotL <c>.(M[b\<turnstile>c>a]) z" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4085
  finally show "(NotL <c>.M z){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotL <c>.M z)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4086
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4088
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "d\<sharp>b" "d\<sharp>a" "d\<sharp>x" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4089
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4090
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4091
  show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4092
  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
  4093
    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
  4094
    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
  4095
    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
  4096
      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
  4097
    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
  4098
               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
  4099
      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
  4100
    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
  4101
      using new by (simp add: fresh_fun_simp_AndR fresh_prod)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4102
    also have "\<dots> \<longrightarrow>\<^sub>a* (AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e')[e'\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4103
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4104
      by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4105
    also have "\<dots> = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4106
      by (auto simp: fresh_prod fresh_atm subst_fresh crename_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4107
    also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) a" using ih1 ih2 by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4108
    also have "\<dots> = (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4109
    finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4110
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4112
    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
  4113
    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
  4114
      using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4115
    also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) e" using ih1 ih2 by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4116
    finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4117
      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
  4118
  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
  4119
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4121
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4122
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4123
  have "(AndL1 (u).M v){b:=(x).Ax x a} = AndL1 (u).(M{b:=(x).Ax x a}) v" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4124
  also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4125
  finally show "(AndL1 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4126
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4128
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4129
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4130
  have "(AndL2 (u).M v){b:=(x).Ax x a} = AndL2 (u).(M{b:=(x).Ax x a}) v" using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4131
  also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4132
  finally show "(AndL2 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4133
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4135
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4136
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4137
  show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4138
  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
  4139
    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
  4140
    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
  4141
    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
  4142
      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
  4143
    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
  4144
             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
  4145
    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
  4146
      using new by (simp add: fresh_fun_simp_OrR1)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4147
    also have "\<dots> \<longrightarrow>\<^sub>a* (OrR1 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4148
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4149
      by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4150
    also have "\<dots> = OrR1 <c>.M{b:=(x).Ax x a} a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4151
      by (auto simp: fresh_prod fresh_atm crename_fresh subst_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4152
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4153
    also have "\<dots> = (OrR1 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4154
    finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4155
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4157
    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
  4158
    have "(OrR1 <c>.M d){b:=(x).Ax x a} = OrR1 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4159
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4160
    finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4161
  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
  4162
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4164
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4165
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4166
  show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4167
  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
  4168
    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
  4169
    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
  4170
    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
  4171
      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
  4172
    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
  4173
             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
  4174
    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
  4175
      using new by (simp add: fresh_fun_simp_OrR2)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4176
    also have "\<dots> \<longrightarrow>\<^sub>a* (OrR2 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4177
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4178
      by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4179
    also have "\<dots> = OrR2 <c>.M{b:=(x).Ax x a} a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4180
      by (auto simp: fresh_prod fresh_atm crename_fresh subst_fresh)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4181
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4182
    also have "\<dots> = (OrR2 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4183
    finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4184
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4186
    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
  4187
    have "(OrR2 <c>.M d){b:=(x).Ax x a} = OrR2 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4188
    also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4189
    finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4190
  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
  4191
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4193
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "v\<sharp>b" "v\<sharp>a" "v\<sharp>x" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4194
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4195
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4196
  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
  4197
    using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4198
  also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[b\<turnstile>c>a]) (v).(N[b\<turnstile>c>a]) z" using ih1 ih2 by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4199
  finally show "(OrL (u).M (v).N z){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[b\<turnstile>c>a]" using fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4200
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4202
  have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "z\<sharp>d" "c\<sharp>d" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4203
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4204
  show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4205
  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
  4206
    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
  4207
    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
  4208
    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
  4209
      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
  4210
    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
  4211
                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
  4212
    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
  4213
      using new by (simp add: fresh_fun_simp_ImpR)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4214
    also have "\<dots> \<longrightarrow>\<^sub>a* (ImpR z.<c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4215
      using new 
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4216
      by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4217
    also have "\<dots> = ImpR z.<c>.M{b:=(x).Ax x a} a" using fs new
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4218
      by (auto simp: fresh_prod crename_fresh subst_fresh fresh_atm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4219
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpR z.<c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4220
    also have "\<dots> = (ImpR z.<c>.M b)[b\<turnstile>c>a]" using eq fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4221
    finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using eq by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4222
  next
9be4ab2acc13 split Class.thy into parts to 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
    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
  4224
    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
  4225
    have "(ImpR (z).<c>.M d){b:=(x).Ax x a} = ImpR (z).<c>.(M{b:=(x).Ax x a}) d" using fs neq by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4226
    also have "\<dots> \<longrightarrow>\<^sub>a* ImpR (z).<c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4227
    finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using neq fs by simp
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4228
  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
  4229
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4231
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4232
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4233
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4234
  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
  4235
    using fs by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4236
  also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[b\<turnstile>c>a]) (u).(N[b\<turnstile>c>a]) v" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4237
    using ih1 ih2 by (auto intro: a_star_congs)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48567
diff changeset
  4238
  finally show "(ImpL <c>.M (u).N v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[b\<turnstile>c>a]" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4239
    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
  4240
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
  4241
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61594
diff changeset
  4242
text \<open>substitution lemmas\<close>
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4243
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4244
lemma not_Ax1: "\<not>(b\<sharp>M) \<Longrightarrow> M{b:=(y).Q} \<noteq> Ax x a"
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4245
proof (nominal_induct M avoiding: b y Q x a rule: trm.strong_induct)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4246
  case (NotR name trm coname)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4247
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4248
    by (metis fin.intros(1) fin_rest_elims(2) subst_not_fin2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4249
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4250
  case (AndR coname1 trm1 coname2 trm2 coname3)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4251
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4252
    by (metis fin.intros(1) fin_rest_elims(3) subst_not_fin2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4253
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4254
  case (OrR1 coname1 trm coname2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4255
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4256
    by (metis fin.intros(1) fin_rest_elims(4) subst_not_fin2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4257
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4258
  case (OrR2 coname1 trm coname2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4259
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4260
    by (metis fin.intros(1) fin_rest_elims(5) subst_not_fin2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4261
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4262
  case (ImpR name coname1 trm coname2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4263
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4264
    by (metis fin.intros(1) fin_rest_elims(6) subst_not_fin2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4265
qed (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4266
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4267
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4268
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4269
lemma not_Ax2: "\<not>(x\<sharp>M) \<Longrightarrow> M{x:=<b>.Q} \<noteq> Ax y a"
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4270
proof (nominal_induct M avoiding: b y Q x a rule: trm.strong_induct)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4271
  case (NotL coname trm name)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4272
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4273
    by (metis fic.intros(1) fic_rest_elims(2) not_fic_subst1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4274
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4275
  case (AndL1 name1 trm name2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4276
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4277
    by (metis fic.intros(1) fic_rest_elims(4) not_fic_subst1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4278
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4279
  case (AndL2 name1 trm name2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4280
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4281
    by (metis fic.intros(1) fic_rest_elims(5) not_fic_subst1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4282
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4283
  case (OrL name1 trm1 name2 trm2 name3)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4284
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4285
    by (metis fic.intros(1) fic_rest_elims(3) not_fic_subst1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4286
next
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4287
  case (ImpL coname trm1 name1 trm2 name2)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4288
  then show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4289
    by (metis fic.intros(1) fic_rest_elims(6) not_fic_subst1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4290
qed (auto simp: fresh_atm abs_fresh abs_supp fin_supp)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4291
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
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
  4294
  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
  4295
  shows "N{y:=<c>.P}{x:=<c>.P} = N{x:=<c>.Ax y c}{y:=<c>.P}"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4296
  using a
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4297
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
  4298
  case (Cut d M u M' x' y' c P)
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  4299
  with assms show ?case
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4300
    apply (simp add: abs_fresh)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4301
    by (smt (verit) forget(1) not_Ax2)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL d M u)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4304
  obtain x'::name and x''::name 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4305
    where "x' \<sharp> (P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},y,x)"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4306
      and  "x''\<sharp> (P,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},Ax y c,y,x)"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4307
    by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4308
  then show ?case
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4309
    using NotL
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4310
    by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4311
        fresh_atm abs_fresh fresh_fun_simp_NotL fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4312
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 u M d)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4314
  obtain x'::name and x''::name 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4315
    where "x' \<sharp> (P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4316
      and  "x''\<sharp> (P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4317
    by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4318
  then show ?case
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4319
    using AndL1
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4320
    by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4321
        fresh_atm abs_fresh fresh_fun_simp_AndL1 fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4322
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL2 u M d)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4324
  obtain x'::name and x''::name 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4325
    where "x' \<sharp> (P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4326
      and  "x''\<sharp> (P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4327
    by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4328
  then show ?case
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4329
    using AndL2
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4330
    by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4331
        fresh_atm abs_fresh fresh_fun_simp_AndL2 fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4332
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL x1 M x2 M' x3)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4334
  obtain x'::name and x''::name 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4335
    where "x' \<sharp> (P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4336
                       M'{y:=<c>.P},M'{x:=<c>.Ax y c}{y:=<c>.P},x1,x2,x3,y,x)"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4337
      and  "x''\<sharp> (P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4338
                                        M'{x:=<c>.Ax y c},M'{x:=<c>.Ax y c}{y:=<c>.P},x1,x2,x3,y,x)"
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4339
    by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4340
  then show ?case
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4341
    using OrL
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4342
    by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4343
        fresh_atm abs_fresh fresh_fun_simp_OrL fresh_prod 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
  4344
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL a M x1 M' x2)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4346
  obtain x'::name and x''::name 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4347
    where "x' \<sharp> (P,M{x2:=<c>.P},M{x:=<c>.Ax x2 c}{x2:=<c>.P},
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4348
                                        M'{x2:=<c>.P},M'{x:=<c>.Ax x2 c}{x2:=<c>.P},x1,y,x)"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4349
      and  "x''\<sharp> (P,Ax y c,M{x2:=<c>.Ax y c},M{x2:=<c>.Ax y c}{y:=<c>.P},
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4350
                                        M'{x2:=<c>.Ax y c},M'{x2:=<c>.Ax y c}{y:=<c>.P},x1,x2,y,x)"
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4351
    by (meson exists_fresh(1) fs_name1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4352
  then show ?case
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4353
    using ImpL
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4354
    by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4355
        fresh_atm abs_fresh fresh_fun_simp_ImpL fresh_prod subst_fresh)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4356
qed (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4357
9be4ab2acc13 split Class.thy into parts to 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
lemma interesting_subst1':
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4359
  assumes "x\<noteq>y" "x\<sharp>P" "y\<sharp>P" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4360
  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
  4361
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
  4362
  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
  4363
  proof (cases "c=a")
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4364
    case True with assms show ?thesis 
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4365
      by (simp add: interesting_subst1)
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4367
    case False 
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4368
    then have "N{x:=<a>.Ax y a} = N{x:=<c>.([(c,a)]\<bullet>Ax y a)}"
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4369
      by (simp add: fresh_atm(2,4) fresh_prod substn_rename4)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4370
    with assms show ?thesis
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4371
      by (simp add: interesting_subst1 swap_simps) 
36277
9be4ab2acc13 split Class.thy into parts to 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
  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
  4373
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
  4374
9be4ab2acc13 split Class.thy into parts to 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
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
  4376
  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
  4377
  shows "N{a:=(y).P}{b:=(y).P} = N{b:=(y).Ax y a}{a:=(y).P}"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4378
  using a
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4379
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
  4380
  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
  4381
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4382
    by (auto simp: abs_fresh fresh_atm forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4383
next 
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut d M u M' x' y' c P)
41893
dde7df1176b7 eliminated prems;
wenzelm
parents: 41798
diff changeset
  4385
  with assms show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4386
    apply(auto simp: trm.inject)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4387
      apply (force simp add: abs_fresh forget)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4388
     apply (simp add: abs_fresh forget)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4389
    by (smt (verit, ccfv_threshold) abs_fresh(1) better_Cut_substc forget(2) fresh_prod not_Ax1)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4392
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4393
    by (auto simp: abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4394
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotR u M d)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4396
  obtain a' a''::coname 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4397
    where "a' \<sharp> (b,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P},u,y)"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4398
      and  "a'' \<sharp> (P,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P},Ax y a,y,d)"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4399
    by (meson exists_fresh'(2) fs_coname1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4400
  with NotR show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4401
    by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_NotR)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR d1 M d2 M' d3)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4404
  obtain a' a''::coname 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4405
    where "a' \<sharp> (P,M{d3:=(y).P},M{b:=(y).Ax y d3}{d3:=(y).P},
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4406
                   M'{d3:=(y).P},M'{b:=(y).Ax y d3}{d3:=(y).P},d1,d2,d3,b,y)"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4407
      and "a'' \<sharp> (P,Ax y a,M{d3:=(y).Ax y a},M{d3:=(y).Ax y a}{a:=(y).P},
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4408
                   M'{d3:=(y).Ax y a},M'{d3:=(y).Ax y a}{a:=(y).P},d1,d2,d3,y,b)"
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4409
    by (meson exists_fresh'(2) fs_coname1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4410
  with AndR show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4411
    apply(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4412
     apply(auto simp only: fresh_prod fresh_fun_simp_AndR)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4413
     apply (force simp:  forget abs_fresh subst_fresh fresh_atm)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4414
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4417
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4418
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4419
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4421
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4422
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4423
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR1 d M e)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4425
  obtain a' a''::coname 
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4426
    where "a' \<sharp> (b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4427
      and "a'' \<sharp> (b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4428
    by (meson exists_fresh'(2) fs_coname1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4429
  with OrR1 show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4430
    by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_OrR1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4431
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR2 d M e)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4433
  obtain a' a''::coname 
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4434
    where "a' \<sharp> (b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4435
      and  "a'' \<sharp> (b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4436
    by (meson exists_fresh'(2) fs_coname1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4437
  with OrR2 show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4438
    by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_OrR2)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4439
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4441
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4442
    by(auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4443
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4445
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4446
    by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4447
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpR u e M d)
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4449
  obtain a' a''::coname 
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4450
    where "a' \<sharp> (b,e,d,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P})"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4451
      and  "a'' \<sharp> (e,d,P,Ax y a,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P})"
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4452
    by (meson exists_fresh'(2) fs_coname1)
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4453
  with ImpR show ?case
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4454
    by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_ImpR)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4456
9be4ab2acc13 split Class.thy into parts to 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
lemma interesting_subst2':
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4458
  assumes "a\<noteq>b" "a\<sharp>P" "b\<sharp>P" 
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4459
  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
  4460
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
  4461
  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
  4462
  proof (cases "z=y")
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4463
    case True then show ?thesis using assms by (simp add: interesting_subst2)
36277
9be4ab2acc13 split Class.thy into parts to 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
  next
80618
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4465
    case False 
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4466
    then have "N{b:=(z).Ax z a} = N{b:=(y).([(y,z)]\<bullet>Ax z a)}"
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4467
      by (metis fresh_atm(1,3) fresh_prod substc_rename2 trm.fresh(1))
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4468
    with False assms show ?thesis
7157c61c8461 More simplification of apply proofs
paulson <lp15@cam.ac.uk>
parents: 80614
diff changeset
  4469
      by (simp add: interesting_subst2 calc_atm)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4470
  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
  4471
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
  4472
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4473
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4475
  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
  4476
  shows "M{x:=<a>.P}{b:=(y).Q} = M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4477
  using a
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4478
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
  4479
  case (Ax z c)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4480
  then show ?case
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4481
    by (auto simp add: abs_fresh fresh_prod fresh_atm subst_fresh trm.inject 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
  4482
next
9be4ab2acc13 split Class.thy into parts to 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
  case (Cut c M z N)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4484
  then show ?case
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4485
    apply (clarsimp simp add: abs_fresh fresh_prod fresh_atm subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4486
    apply (metis forget not_Ax1 not_Ax2)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4487
    done
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4488
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4490
  then show ?case
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4491
    by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_NotR)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL c M z)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4494
  obtain x'::name where "x' \<sharp> (P,M{x:=<a>.P},P{b:=(y).Q},M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},y,Q)"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4495
    by (meson exists_fresh(1) fs_name1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4496
  with NotL show ?case  
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4497
    apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh subst_fresh fresh_fun_simp_NotL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4498
    by (metis fresh_fun_simp_NotL fresh_prod subst_fresh(5))
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4501
  then show ?case  
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4502
    by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_AndR)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 z1 M z2)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4505
  obtain x'::name where "x' \<sharp> (P,M{x:=<a>.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}})"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4506
    by (meson exists_fresh(1) fs_name1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4507
  with AndL1 show ?case  
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4508
    apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_AndL1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4509
    by (metis fresh_atm(1) fresh_fun_simp_AndL1 fresh_prod subst_fresh(5))
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL2 z1 M z2)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4512
  obtain x'::name where "x' \<sharp> (P,M{x:=<a>.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}})"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4513
    by (meson exists_fresh(1) fs_name1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4514
  with AndL2 show ?case  
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4515
    apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_AndL2)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4516
    by (metis fresh_atm(1) fresh_fun_simp_AndL2 fresh_prod subst_fresh(5))
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL z1 M z2 N z3)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4519
  obtain x'::name where "x' \<sharp> (P,M{x:=<a>.P},M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},z2,z3,a,y,Q,
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4520
                                     P{b:=(y).Q},N{x:=<a>.P},N{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},z1)"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4521
    by (meson exists_fresh(1) fs_name1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4522
  with OrL show ?case  
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4523
    apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_OrL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4524
    by (metis (full_types) fresh_atm(1) fresh_fun_simp_OrL fresh_prod subst_fresh(5))
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4527
  then show ?case  
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4528
    by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_OrR1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4529
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4531
  then show ?case  
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4532
    by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_OrR2)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4533
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4535
  then show ?case  
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4536
    by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_ImpR)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL c M z N u)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4539
  obtain z'::name where "z' \<sharp> (P,P{b:=(y).Q},M{u:=<a>.P},N{u:=<a>.P},y,Q,
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4540
                        M{b:=(y).Q}{u:=<a>.P{b:=(y).Q}},N{b:=(y).Q}{u:=<a>.P{b:=(y).Q}},z)"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4541
    by (meson exists_fresh(1) fs_name1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4542
  with ImpL show ?case  
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4543
    apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_ImpL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4544
    by (metis (full_types) fresh_atm(1) fresh_prod subst_fresh(5) fresh_fun_simp_ImpL)
36277
9be4ab2acc13 split Class.thy into parts to 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
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
  4546
9be4ab2acc13 split Class.thy into parts to 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
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
  4548
  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
  4549
  shows "M{a:=(x).N}{y:=<b>.P} = M{y:=<b>.P}{a:=(x).N{y:=<b>.P}}"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4550
  using a
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4551
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
  4552
  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
  4553
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4554
    by (auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4555
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4557
  then show ?case
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4558
    apply (clarsimp simp add: abs_fresh fresh_prod fresh_atm subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4559
    apply (metis forget not_Ax1 not_Ax2)
36277
9be4ab2acc13 split Class.thy into parts to 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
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotR z M' d) 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4563
  obtain a'::coname where "a' \<sharp> (y,P,N,N{y:=<b>.P},M'{d:=(x).N},M'{y:=<b>.P}{d:=(x).N{y:=<b>.P}})"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4564
    by (meson exists_fresh'(2) fs_coname1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4565
  with NotR show ?case
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4566
    apply(simp add: fresh_atm subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4567
    apply(auto simp only: fresh_prod fresh_fun_simp_NotR; simp add: abs_fresh NotR)
36277
9be4ab2acc13 split Class.thy into parts to 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
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (NotL d M' z) 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4571
  obtain x'::name where "x' \<sharp> (z,y,P,N,N{y:=<b>.P},M'{y:=<b>.P},M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}})"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4572
    by (meson exists_fresh(1) fs_name1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4573
  with NotL show ?case  
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4574
    apply(simp add: fresh_prod fresh_atm abs_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4575
    apply(auto simp only: fresh_fun_simp_NotL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4576
    by (auto simp: subst_fresh abs_fresh fresh_atm forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4577
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndR d M' e M'' f) 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4579
  obtain a'::coname where "a' \<sharp> (P,b,d,e,N,N{y:=<b>.P},M'{f:=(x).N},M''{f:=(x).N},
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4580
                  M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}},M''{y:=<b>.P}{f:=(x).N{y:=<b>.P}})"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4581
    by (meson exists_fresh'(2) fs_coname1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4582
  with AndR show ?case
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4583
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4584
    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
  4585
                  M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}},M''{y:=<b>.P}{f:=(x).N{y:=<b>.P}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4586
     apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4587
     apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4588
     apply(simp add: fresh_fun_simp_AndR)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4589
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4590
    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
  4591
    done
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4592
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL1 z M' u) 
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4594
  obtain x'::name where "x' \<sharp> (P,b,z,u,x,N,M'{y:=<b>.P},M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}})"
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4595
    by (meson exists_fresh(1) fs_name1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4596
  with AndL1 show ?case  
80651
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4597
    by (force simp add: fresh_prod fresh_atm fresh_fun_simp_AndL1 subst_fresh abs_fresh 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
  4598
next
9be4ab2acc13 split Class.thy into parts to 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
  case (AndL2 z M' u) 
80651
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4600
  obtain x'::name where "x' \<sharp> (P,b,z,u,x,N,M'{y:=<b>.P},M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}})"
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4601
    by (meson exists_fresh(1) fs_name1)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4602
  with AndL2 show ?case  
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4603
    by (force simp add: fresh_prod fresh_atm fresh_fun_simp_AndL2 subst_fresh abs_fresh 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
  4604
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrL u M' v M'' w) 
80651
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4606
  obtain x'::name where "x' \<sharp> (P,b,u,w,v,N,N{y:=<b>.P},M'{y:=<b>.P},M''{y:=<b>.P},
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4607
                  M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}},M''{y:=<b>.P}{a:=(x).N{y:=<b>.P}})"
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4608
    by (meson exists_fresh(1) fs_name1)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4609
  with OrL show ?case  
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4610
    by (force simp add: fresh_prod fresh_atm fresh_fun_simp_OrL subst_fresh abs_fresh 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
  4611
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR1 e M' f) 
80651
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4613
  obtain c'::coname where c': "c' \<sharp> (P,b,e,f,x,N,N{y:=<b>.P},
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4614
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}})"
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4615
    by (meson exists_fresh(2) fs_coname1)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4616
  with OrR1 show ?case  
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4617
    apply (clarsimp simp: fresh_prod fresh_fun_simp_OrR1)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4618
    apply (clarsimp simp: subst_fresh abs_fresh fresh_atm)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4619
    using c' apply (auto simp: fresh_prod fresh_fun_simp_OrR1)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4620
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (OrR2 e M' f) 
80651
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4623
  obtain c'::coname where c': "c' \<sharp> (P,b,e,f,x,N,N{y:=<b>.P},
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4624
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}})"
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4625
    by (meson exists_fresh(2) fs_coname1)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4626
  with OrR2 show ?case  
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4627
    apply (clarsimp simp: fresh_prod fresh_fun_simp_OrR2)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4628
    apply (clarsimp simp: subst_fresh abs_fresh fresh_atm)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4629
    using c' apply (auto simp: fresh_prod fresh_fun_simp_OrR2)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4630
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpR x e M' f) 
80651
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4633
  obtain c'::coname where c': "c' \<sharp> (P,b,e,f,x,N,N{y:=<b>.P},
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4634
                                     M'{f:=(x).N},M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}})"
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4635
    by (meson exists_fresh(2) fs_coname1)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4636
  with ImpR show ?case  
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4637
    apply (clarsimp simp: fresh_prod fresh_fun_simp_ImpR)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4638
    apply (clarsimp simp: subst_fresh abs_fresh fresh_atm)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4639
    using c' apply (auto simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_fun_simp_ImpR)
36277
9be4ab2acc13 split Class.thy into parts to 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
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  case (ImpL e M' v M'' w) 
80651
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4643
  obtain z'::name where z': "z' \<sharp> (P,b,e,w,v,N,N{y:=<b>.P},M'{w:=<b>.P},M''{w:=<b>.P},
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4644
                  M'{w:=<b>.P}{a:=(x).N{w:=<b>.P}},M''{w:=<b>.P}{a:=(x).N{w:=<b>.P}})"
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4645
    by (meson exists_fresh(1) fs_name1)
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4646
  with ImpL show ?case  
2cffa664482d tidied more apply proofs
paulson <lp15@cam.ac.uk>
parents: 80620
diff changeset
  4647
    by (force simp add: fresh_prod fresh_atm fresh_fun_simp_ImpL subst_fresh abs_fresh 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
  4648
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
  4649
9be4ab2acc13 split Class.thy into parts to 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
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
  4651
  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
  4652
  shows "N{x:=<a>.M}{y:=<c>.P} = N{y:=<c>.P}{x:=<a>.(M{y:=<c>.P})}"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4653
  using a
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4654
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
  4655
  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
  4656
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4657
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4658
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4660
  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
  4661
    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
  4662
    apply(auto)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4663
       apply(auto simp add: fresh_atm trm.inject forget)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4664
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4665
    apply (metis forget(1) not_Ax2)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4666
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4667
     apply(rule sym)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4668
     apply(rule trans)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4669
      apply(rule better_Cut_substn)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4670
       apply(simp add: abs_fresh subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4671
      apply(simp add: fresh_prod subst_fresh fresh_atm)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4672
     apply(simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4673
    by (metis forget(1) not_Ax2)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4676
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4677
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4678
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4680
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4681
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4682
     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}})")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4683
      apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4684
      apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4685
      apply(simp add: fresh_fun_simp_NotL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4686
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4687
     apply(rule exists_fresh'(1)[OF fs_name1])
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4688
    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}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4689
     apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4690
     apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4691
     apply(simp add: fresh_fun_simp_NotL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4692
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4693
    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
  4694
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4697
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4698
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4699
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4701
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4702
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4703
     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}})")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4704
      apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4705
      apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4706
      apply(simp add: fresh_fun_simp_AndL1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4707
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4708
     apply(rule exists_fresh'(1)[OF fs_name1])
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4709
    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}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4710
     apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4711
     apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4712
     apply(simp add: fresh_fun_simp_AndL1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4713
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4714
    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
  4715
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4718
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4719
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4720
     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}})")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4721
      apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4722
      apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4723
      apply(simp add: fresh_fun_simp_AndL2)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4724
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4725
     apply(rule exists_fresh'(1)[OF fs_name1])
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4726
    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}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4727
     apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4728
     apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4729
     apply(simp add: fresh_fun_simp_AndL2)     
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4730
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4731
    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
  4732
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4735
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4736
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4737
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4739
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4740
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4741
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4743
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4744
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4745
     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}},
36277
9be4ab2acc13 split Class.thy into parts to 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
                                      x1,x2,x3,M''{x:=<a>.M},M''{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4747
      apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4748
      apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4749
      apply(simp add: fresh_fun_simp_OrL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4750
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4751
     apply(rule exists_fresh'(1)[OF fs_name1])
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4752
    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
  4753
                                      x1,x2,x3,M''{y:=<c>.P},M''{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4754
     apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4755
     apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4756
     apply(simp add: fresh_fun_simp_OrL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4757
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4758
    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
  4759
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4762
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4763
    by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4764
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4766
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4767
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4768
     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}},
36277
9be4ab2acc13 split Class.thy into parts to 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
                                      x1,x2,M''{x2:=<a>.M},M''{y:=<c>.P}{x2:=<a>.M{y:=<c>.P}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4770
      apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4771
      apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4772
      apply(simp add: fresh_fun_simp_ImpL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4773
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4774
     apply(rule exists_fresh'(1)[OF fs_name1])
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4775
    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
  4776
                                      x1,x2,M''{x2:=<c>.P},M''{x2:=<c>.P}{x:=<a>.M{x2:=<c>.P}})")
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4777
     apply(erule exE, simp add: fresh_prod)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4778
     apply(erule conjE)+
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4779
     apply(simp add: fresh_fun_simp_ImpL)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4780
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4781
    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
  4782
    done
9be4ab2acc13 split Class.thy into parts to 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
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
  4784
9be4ab2acc13 split Class.thy into parts to 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
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
  4786
  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
  4787
  shows "N{a:=(x).M}{c:=(y).P} = N{c:=(y).P}{a:=(x).(M{c:=(y).P})}"
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4788
  using a
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4789
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
  4790
  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
  4791
  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
  4792
    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
  4793
    apply(auto)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4794
       apply(simp add: fresh_atm)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4795
       apply(simp add: trm.inject)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4796
       apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4797
      apply (metis forget(2) not_Ax1)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4798
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4799
    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
  4800
    apply(rule trans)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4801
     apply(rule better_Cut_substc)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4802
      apply(simp add: fresh_prod subst_fresh fresh_atm)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4803
     apply(simp add: abs_fresh subst_fresh)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4804
    apply(auto)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4805
    by (metis forget(2) not_Ax1)
36277
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4808
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4809
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4810
     apply(generate_fresh "coname")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4811
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4812
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4813
      apply(simp add: abs_fresh subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4814
      apply(simp add: abs_fresh subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4815
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4816
    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
  4817
    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
  4818
    apply(fresh_fun_simp)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4819
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4820
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4823
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4824
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4825
     apply(generate_fresh "coname")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4826
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4827
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4828
      apply(simp add: abs_fresh subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4829
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4830
    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
  4831
    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
  4832
    apply(fresh_fun_simp)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4833
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4834
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4837
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4838
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4839
     apply(generate_fresh "coname")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4840
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4841
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4842
      apply(simp add: abs_fresh subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4843
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4844
    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
  4845
    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
  4846
    apply(fresh_fun_simp)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4847
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4848
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4851
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4852
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4853
     apply(generate_fresh "coname")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4854
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4855
     apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4856
      apply(simp add: abs_fresh subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4857
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4858
    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
  4859
    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
  4860
    apply(fresh_fun_simp)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4861
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4862
    done
9be4ab2acc13 split Class.thy into parts to 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
next
9be4ab2acc13 split Class.thy into parts to 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
  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
  4865
  then show ?case
80138
a30a1385f7d0 Starting to tidy HOL-Nominal-Examples
paulson <lp15@cam.ac.uk>
parents: 74101
diff changeset
  4866
    apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4867
      apply(generate_fresh "coname")
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4868
      apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4869
      apply(fresh_fun_simp)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4870
      apply(simp add: abs_fresh subst_fresh)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4871
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4872
    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
  4873
    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
  4874
    apply(fresh_fun_simp)
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4875
     apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)+
36277
9be4ab2acc13 split Class.thy into parts to conserve a bit of memory and increase the chance of making it work on Cygwin with only 2 GB available;
wenzelm
parents:
diff changeset
  4876
    done
80620
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4877
qed (auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget)
ee77d9d40be6 More simplification of a nominal example
paulson <lp15@cam.ac.uk>
parents: 80618
diff changeset
  4878
36277
9be4ab2acc13 split Class.thy into parts to 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
9be4ab2acc13 split Class.thy into parts to 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
end